(**************************************************************************) (* *) (* Landau's "Grundlagen der Analysis", formal specification for AUTOMATH *) (* Copyright (C) 1977, L.S. van Benthem Jutting *) (* 1992, revised by F. Wiedijk *) (* 2008, 2014, revised by F. Guidi *) (* *) (* Mechanized translation for COQ version 8 by F. Guidi *) (* *) (**************************************************************************) (* This file was generated by Helena 0.8.3 M (June 2015): do not edit *****) (* constant 1 *) Definition l_imp : (forall (a:Prop), (forall (b:Prop), Prop)). exact (fun (a:Prop) => (fun (b:Prop) => (forall (x:a), b))). Time Defined. (* constant 2 *) Definition l_mp : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (i:l_imp a b), b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => (fun (i:l_imp a b) => i a1)))). Time Defined. (* constant 3 *) Definition l_refimp : (forall (a:Prop), l_imp a a). exact (fun (a:Prop) => (fun (x:a) => x)). Time Defined. (* constant 4 *) Definition l_trimp : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_imp a b), (forall (j:l_imp b c), l_imp a c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_imp a b) => (fun (j:l_imp b c) => (fun (x:a) => j (i x))))))). Time Defined. (* constant 5 *) Axiom l_con : Prop. (* constant 6 *) Definition l_not : (forall (a:Prop), Prop). exact (fun (a:Prop) => l_imp a l_con). Time Defined. (* constant 7 *) Definition l_wel : (forall (a:Prop), Prop). exact (fun (a:Prop) => l_not (l_not a)). Time Defined. (* constant 8 *) Definition l_weli : (forall (a:Prop), (forall (a1:a), l_wel a)). exact (fun (a:Prop) => (fun (a1:a) => (fun (x:l_not a) => x a1))). Time Defined. (* constant 9 *) Axiom l_et : (forall (a:Prop), (forall (w:l_wel a), a)). (* constant 10 *) Definition l_cone : (forall (a:Prop), (forall (c1:l_con), a)). exact (fun (a:Prop) => (fun (c1:l_con) => l_et a (fun (x:l_not a) => c1))). Time Defined. (* constant 11 *) Definition l_imp_th1 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp a b), (forall (j:l_imp (l_not a) b), b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp a b) => (fun (j:l_imp (l_not a) b) => l_et b (fun (x:l_not b) => x (j (l_trimp a b l_con i x))))))). Time Defined. (* constant 12 *) Definition l_imp_th2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), l_imp a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => l_trimp a l_con b n (fun (x:l_con) => l_cone b x)))). Time Defined. (* constant 13 *) Definition l_imp_th3 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not b), (forall (i:l_imp a b), l_not a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not b) => (fun (i:l_imp a b) => l_trimp a b l_con i n)))). Time Defined. (* constant 14 *) Definition l_imp_th4 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (n:l_not b), l_not (l_imp a b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => (fun (n:l_not b) => (fun (x:l_imp a b) => l_imp_th3 a b n x a1))))). Time Defined. (* constant 15 *) Definition l_imp_th5 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_imp a b)), a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_imp a b)) => l_et a (fun (x:l_not a) => n (l_imp_th2 a b x))))). Time Defined. (* constant 16 *) Definition l_imp_th6 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_imp a b)), l_not b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_imp a b)) => (fun (x:b) => n (fun (y:a) => x))))). Time Defined. (* constant 17 *) Definition l_imp_th7 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not b), (forall (i:l_imp (l_not a) b), a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not b) => (fun (i:l_imp (l_not a) b) => l_et a (l_imp_th3 (l_not a) b n i))))). Time Defined. (* constant 18 *) Definition l_cp : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp (l_not b) (l_not a)), l_imp a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp (l_not b) (l_not a)) => (fun (x:a) => l_imp_th7 b (l_not a) (l_weli a x) i)))). Time Defined. (* constant 19 *) Definition l_obvious : Prop. exact (l_imp l_con l_con). Time Defined. (* constant 20 *) Definition l_obviousi : l_obvious. exact (l_refimp l_con). Time Defined. (* constant 21 *) Definition l_ec : (forall (a:Prop), (forall (b:Prop), Prop)). exact (fun (a:Prop) => (fun (b:Prop) => l_imp a (l_not b))). Time Defined. (* constant 22 *) Definition l_eci1 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), l_ec a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => l_imp_th2 a (l_not b) n))). Time Defined. (* constant 23 *) Definition l_eci2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not b), l_ec a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not b) => (fun (x:a) => n)))). Time Defined. (* constant 24 *) Definition l_ec_th1 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp a (l_not b)), l_ec a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp a (l_not b)) => i))). Time Defined. (* constant 25 *) Definition l_ec_th2 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp b (l_not a)), l_ec a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp b (l_not a)) => (fun (x:a) => (fun (y:b) => i y x))))). Time Defined. (* constant 26 *) Definition l_comec : (forall (a:Prop), (forall (b:Prop), (forall (e:l_ec a b), l_ec b a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (e:l_ec a b) => l_ec_th2 b a e))). Time Defined. (* constant 27 *) Definition l_ece1 : (forall (a:Prop), (forall (b:Prop), (forall (e:l_ec a b), (forall (a1:a), l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (e:l_ec a b) => (fun (a1:a) => e a1)))). Time Defined. (* constant 28 *) Definition l_ece2 : (forall (a:Prop), (forall (b:Prop), (forall (e:l_ec a b), (forall (b1:b), l_not a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (e:l_ec a b) => (fun (b1:b) => l_imp_th3 a (l_not b) (l_weli b b1) e)))). Time Defined. (* constant 29 *) Definition l_ec_th3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec a b), (forall (i:l_imp c a), l_ec c b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec a b) => (fun (i:l_imp c a) => l_trimp c a (l_not b) i e))))). Time Defined. (* constant 30 *) Definition l_ec_th4 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec a b), (forall (i:l_imp c b), l_ec a c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec a b) => (fun (i:l_imp c b) => l_comec c a (l_ec_th3 b a c (l_comec a b e) i)))))). Time Defined. (* constant 31 *) Definition l_and : (forall (a:Prop), (forall (b:Prop), Prop)). exact (fun (a:Prop) => (fun (b:Prop) => l_not (l_ec a b))). Time Defined. (* constant 32 *) Definition l_andi : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (b1:b), l_and a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => (fun (b1:b) => l_imp_th4 a (l_not b) a1 (l_weli b b1))))). Time Defined. (* constant 33 *) Definition l_ande1 : (forall (a:Prop), (forall (b:Prop), (forall (a1:l_and a b), a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:l_and a b) => l_imp_th5 a (l_not b) a1))). Time Defined. (* constant 34 *) Definition l_ande2 : (forall (a:Prop), (forall (b:Prop), (forall (a1:l_and a b), b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:l_and a b) => l_et b (l_imp_th6 a (l_not b) a1)))). Time Defined. (* constant 35 *) Definition l_comand : (forall (a:Prop), (forall (b:Prop), (forall (a1:l_and a b), l_and b a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:l_and a b) => l_andi b a (l_ande2 a b a1) (l_ande1 a b a1)))). Time Defined. (* constant 36 *) Definition l_and_th1 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), l_not (l_and a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => l_weli (l_ec a b) (l_eci1 a b n)))). Time Defined. (* constant 37 *) Definition l_and_th2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not b), l_not (l_and a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not b) => l_weli (l_ec a b) (l_eci2 a b n)))). Time Defined. (* constant 38 *) Definition l_and_th3 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_and a b)), (forall (a1:a), l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_and a b)) => (fun (a1:a) => l_ece1 a b (l_et (l_ec a b) n) a1)))). Time Defined. (* constant 39 *) Definition l_and_th4 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_and a b)), (forall (b1:b), l_not a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_and a b)) => (fun (b1:b) => l_ece2 a b (l_et (l_ec a b) n) b1)))). Time Defined. (* constant 40 *) Definition l_and_th5 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_and a b)), l_not (l_and b a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_and a b)) => l_imp_th3 (l_and b a) (l_and a b) n (fun (x:l_and b a) => l_comand b a x)))). Time Defined. (* constant 41 *) Definition l_and_th6 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and a b), (forall (i:l_imp a c), l_and c b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and a b) => (fun (i:l_imp a c) => l_andi c b (i (l_ande1 a b a1)) (l_ande2 a b a1)))))). Time Defined. (* constant 42 *) Definition l_and_th7 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and a b), (forall (i:l_imp b c), l_and a c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and a b) => (fun (i:l_imp b c) => l_andi a c (l_ande1 a b a1) (i (l_ande2 a b a1))))))). Time Defined. (* constant 43 *) Definition l_or : (forall (a:Prop), (forall (b:Prop), Prop)). exact (fun (a:Prop) => (fun (b:Prop) => l_imp (l_not a) b)). Time Defined. (* constant 44 *) Definition l_ori1 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), l_or a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => l_imp_th2 (l_not a) b (l_weli a a1)))). Time Defined. (* constant 45 *) Definition l_ori2 : (forall (a:Prop), (forall (b:Prop), (forall (b1:b), l_or a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (b1:b) => (fun (x:l_not a) => b1)))). Time Defined. (* constant 46 *) Definition l_or_th1 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp (l_not a) b), l_or a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp (l_not a) b) => i))). Time Defined. (* constant 47 *) Definition l_or_th2 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp (l_not b) a), l_or a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp (l_not b) a) => (fun (x:l_not a) => l_et b (l_imp_th3 (l_not b) a x i))))). Time Defined. (* constant 48 *) Definition l_ore2 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), (forall (n:l_not a), b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => (fun (n:l_not a) => o n)))). Time Defined. (* constant 49 *) Definition l_ore1 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), (forall (n:l_not b), a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => (fun (n:l_not b) => l_et a (l_imp_th3 (l_not a) b n o))))). Time Defined. (* constant 50 *) Definition l_comor : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), l_or b a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => (fun (x:l_not b) => l_ore1 a b o x)))). Time Defined. (* constant 51 *) Definition l_or_th3 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), (forall (m:l_not b), l_not (l_or a b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => (fun (m:l_not b) => l_imp_th4 (l_not a) b n m)))). Time Defined. (* constant 52 *) Definition l_or_th4 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_or a b)), l_not a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_or a b)) => l_imp_th5 (l_not a) b n))). Time Defined. (* constant 53 *) Definition l_or_th5 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_or a b)), l_not b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_or a b)) => l_imp_th6 (l_not a) b n))). Time Defined. (* constant 54 *) Definition l_or_th6 : (forall (a:Prop), l_or a (l_not a)). exact (fun (a:Prop) => l_refimp (l_not a)). Time Defined. (* constant 55 *) Definition l_orapp : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or a b), (forall (i:l_imp a c), (forall (j:l_imp b c), c)))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or a b) => (fun (i:l_imp a c) => (fun (j:l_imp b c) => l_imp_th1 a c i (l_trimp (l_not a) b c o j))))))). Time Defined. (* constant 56 *) Definition l_or_th7 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or a b), (forall (i:l_imp a c), l_or c b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or a b) => (fun (i:l_imp a c) => l_trimp (l_not c) (l_not a) b (fun (x:l_not c) => l_imp_th3 a c x i) o))))). Time Defined. (* constant 57 *) Definition l_or_th8 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or a b), (forall (i:l_imp b c), l_or a c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or a b) => (fun (i:l_imp b c) => l_trimp (l_not a) b c o i))))). Time Defined. (* constant 58 *) Definition l_or_th9 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (d:Prop), (forall (o:l_or a b), (forall (i:l_imp a c), (forall (j:l_imp b d), l_or c d))))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (d:Prop) => (fun (o:l_or a b) => (fun (i:l_imp a c) => (fun (j:l_imp b d) => l_or_th7 a d c (l_or_th8 a b d o j) i))))))). Time Defined. (* constant 59 *) Definition l_or_th10 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), l_imp (l_not a) b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => o))). Time Defined. (* constant 60 *) Definition l_or_th11 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), l_imp (l_not b) a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => l_comor a b o))). Time Defined. (* constant 61 *) Definition l_or_th12 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or (l_not a) b), l_imp a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or (l_not a) b) => l_trimp a (l_wel a) b (fun (x:a) => l_weli a x) o))). Time Defined. (* constant 62 *) Definition l_or_th13 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp a b), l_or (l_not a) b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp a b) => l_trimp (l_wel a) a b (fun (x:l_wel a) => l_et a x) i))). Time Defined. (* constant 63 *) Definition l_or_th14 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or (l_not a) (l_not b)), l_not (l_and a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or (l_not a) (l_not b)) => l_weli (l_ec a b) (l_or_th12 a (l_not b) o)))). Time Defined. (* constant 64 *) Definition l_or_th15 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_and a b)), l_or (l_not a) (l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_and a b)) => l_or_th13 a (l_not b) (l_et (l_ec a b) n)))). Time Defined. (* constant 65 *) Definition l_or_th16 : (forall (a:Prop), (forall (b:Prop), (forall (a1:l_and (l_not a) (l_not b)), l_not (l_or a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:l_and (l_not a) (l_not b)) => l_or_th3 a b (l_ande1 (l_not a) (l_not b) a1) (l_ande2 (l_not a) (l_not b) a1)))). Time Defined. (* constant 66 *) Definition l_or_th17 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_or a b)), l_and (l_not a) (l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_or a b)) => l_andi (l_not a) (l_not b) (l_or_th4 a b n) (l_or_th5 a b n)))). Time Defined. (* constant 67 *) Definition l_orec : (forall (a:Prop), (forall (b:Prop), Prop)). exact (fun (a:Prop) => (fun (b:Prop) => l_and (l_or a b) (l_ec a b))). Time Defined. (* constant 68 *) Definition l_oreci : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), (forall (e:l_ec a b), l_orec a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => (fun (e:l_ec a b) => l_andi (l_or a b) (l_ec a b) o e)))). Time Defined. (* constant 69 *) Definition l_orec_th1 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (n:l_not b), l_orec a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => (fun (n:l_not b) => l_oreci a b (l_ori1 a b a1) (l_eci2 a b n))))). Time Defined. (* constant 70 *) Definition l_orec_th2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), (forall (b1:b), l_orec a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => (fun (b1:b) => l_oreci a b (l_ori2 a b b1) (l_eci1 a b n))))). Time Defined. (* constant 71 *) Definition l_orece1 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_or a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => l_ande1 (l_or a b) (l_ec a b) o))). Time Defined. (* constant 72 *) Definition l_orece2 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_ec a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => l_ande2 (l_or a b) (l_ec a b) o))). Time Defined. (* constant 73 *) Definition l_comorec : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_orec b a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => l_oreci b a (l_comor a b (l_orece1 a b o)) (l_comec a b (l_orece2 a b o))))). Time Defined. (* constant 74 *) Definition l_orec_th3 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), (forall (a1:a), l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => (fun (a1:a) => l_ece1 a b (l_orece2 a b o) a1)))). Time Defined. (* constant 75 *) Definition l_orec_th4 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), (forall (b1:b), l_not a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => (fun (b1:b) => l_ece2 a b (l_orece2 a b o) b1)))). Time Defined. (* constant 76 *) Definition l_orec_th5 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), (forall (n:l_not a), b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => (fun (n:l_not a) => l_ore2 a b (l_orece1 a b o) n)))). Time Defined. (* constant 77 *) Definition l_orec_th6 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), (forall (n:l_not b), a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => (fun (n:l_not b) => l_ore1 a b (l_orece1 a b o) n)))). Time Defined. (* constant 78 *) Definition l_iff : (forall (a:Prop), (forall (b:Prop), Prop)). exact (fun (a:Prop) => (fun (b:Prop) => l_and (l_imp a b) (l_imp b a))). Time Defined. (* constant 79 *) Definition l_iffi : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp a b), (forall (j:l_imp b a), l_iff a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp a b) => (fun (j:l_imp b a) => l_andi (l_imp a b) (l_imp b a) i j)))). Time Defined. (* constant 80 *) Definition l_iff_th1 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (b1:b), l_iff a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => (fun (b1:b) => l_iffi a b (fun (x:a) => b1) (fun (x:b) => a1))))). Time Defined. (* constant 81 *) Definition l_iff_th2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), (forall (m:l_not b), l_iff a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => (fun (m:l_not b) => l_iffi a b (l_imp_th2 a b n) (l_imp_th2 b a m))))). Time Defined. (* constant 82 *) Definition l_iffe1 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_imp a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => l_ande1 (l_imp a b) (l_imp b a) i))). Time Defined. (* constant 83 *) Definition l_iffe2 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_imp b a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => l_ande2 (l_imp a b) (l_imp b a) i))). Time Defined. (* constant 84 *) Definition l_comiff : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_iff b a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => l_iffi b a (l_iffe2 a b i) (l_iffe1 a b i)))). Time Defined. (* constant 85 *) Definition l_iff_th3 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), (forall (a1:a), b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => (fun (a1:a) => l_iffe1 a b i a1)))). Time Defined. (* constant 86 *) Definition l_iff_th4 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), (forall (b1:b), a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => (fun (b1:b) => l_iffe2 a b i b1)))). Time Defined. (* constant 87 *) Definition l_iff_th5 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), (forall (n:l_not a), l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => (fun (n:l_not a) => l_imp_th3 b a n (l_iffe2 a b i))))). Time Defined. (* constant 88 *) Definition l_iff_th6 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), (forall (n:l_not b), l_not a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => (fun (n:l_not b) => l_imp_th3 a b n (l_iffe1 a b i))))). Time Defined. (* constant 89 *) Definition l_iff_th7 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (n:l_not b), l_not (l_iff a b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => (fun (n:l_not b) => l_and_th1 (l_imp a b) (l_imp b a) (l_imp_th4 a b a1 n))))). Time Defined. (* constant 90 *) Definition l_iff_th8 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), (forall (b1:b), l_not (l_iff a b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => (fun (b1:b) => l_and_th2 (l_imp a b) (l_imp b a) (l_imp_th4 b a b1 n))))). Time Defined. (* constant 91 *) Definition l_refiff : (forall (a:Prop), l_iff a a). exact (fun (a:Prop) => l_iffi a a (l_refimp a) (l_refimp a)). Time Defined. (* constant 92 *) Definition l_symiff : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_iff b a))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => l_comiff a b i))). Time Defined. (* constant 93 *) Definition l_triff : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (j:l_iff b c), l_iff a c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (j:l_iff b c) => l_iffi a c (l_trimp a b c (l_iffe1 a b i) (l_iffe1 b c j)) (l_trimp c b a (l_iffe2 b c j) (l_iffe2 a b i))))))). Time Defined. (* constant 94 *) Definition l_iff_th9 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_imp (l_not a) (l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => (fun (x:l_not a) => l_iff_th5 a b i x)))). Time Defined. (* constant 95 *) Definition l_iff_th10 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_imp (l_not b) (l_not a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => (fun (x:l_not b) => l_iff_th6 a b i x)))). Time Defined. (* constant 96 *) Definition l_iff_th11 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_iff (l_not a) (l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => l_iffi (l_not a) (l_not b) (l_iff_th9 a b i) (l_iff_th10 a b i)))). Time Defined. (* constant 97 *) Definition l_iff_th12 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp (l_not a) (l_not b)), (forall (j:l_imp (l_not b) (l_not a)), l_iff a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp (l_not a) (l_not b)) => (fun (j:l_imp (l_not b) (l_not a)) => l_iffi a b (l_cp a b j) (l_cp b a i))))). Time Defined. (* constant 98 *) Definition l_iff_th13 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_iff a (l_not b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => l_iffi a (l_not b) (l_orece2 a b o) (l_comor a b (l_orece1 a b o))))). Time Defined. (* constant 99 *) Definition l_iff_th14 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_iff b (l_not a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => l_iff_th13 b a (l_comorec a b o)))). Time Defined. (* constant 100 *) Definition l_iff_th15 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a (l_not b)), l_orec a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a (l_not b)) => l_oreci a b (l_comor b a (l_iffe2 a (l_not b) i)) (l_iffe1 a (l_not b) i)))). Time Defined. (* constant 101 *) Definition l_iff_th16 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff b (l_not a)), l_orec a b))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff b (l_not a)) => l_comorec b a (l_iff_th15 b a i)))). Time Defined. (* constant 102 *) Definition l_iff_thimp1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (j:l_imp a c), l_imp b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (j:l_imp a c) => l_trimp b a c (l_iffe2 a b i) j))))). Time Defined. (* constant 103 *) Definition l_iff_thimp2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (j:l_imp c a), l_imp c b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (j:l_imp c a) => l_trimp c a b j (l_iffe1 a b i)))))). Time Defined. (* constant 104 *) Definition l_iff_thec1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (e:l_ec a c), l_ec b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (e:l_ec a c) => l_ec_th3 a c b e (l_iffe2 a b i)))))). Time Defined. (* constant 105 *) Definition l_iff_thec2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (e:l_ec c a), l_ec c b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (e:l_ec c a) => l_ec_th4 c a b e (l_iffe2 a b i)))))). Time Defined. (* constant 106 *) Definition l_iff_thand1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (a1:l_and a c), l_and b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (a1:l_and a c) => l_and_th6 a c b a1 (l_iffe1 a b i)))))). Time Defined. (* constant 107 *) Definition l_iff_thand2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (a1:l_and c a), l_and c b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (a1:l_and c a) => l_and_th7 c a b a1 (l_iffe1 a b i)))))). Time Defined. (* constant 108 *) Definition l_iff_thor1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (o:l_or a c), l_or b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (o:l_or a c) => l_or_th7 a c b o (l_iffe1 a b i)))))). Time Defined. (* constant 109 *) Definition l_iff_thor2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (o:l_or c a), l_or c b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (o:l_or c a) => l_or_th8 c a b o (l_iffe1 a b i)))))). Time Defined. (* constant 110 *) Definition l_iff_thorec1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (o:l_orec a c), l_orec b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (o:l_orec a c) => l_oreci b c (l_iff_thor1 a b c i (l_orece1 a c o)) (l_iff_thec1 a b c i (l_orece2 a c o))))))). Time Defined. (* constant 111 *) Definition l_iff_thorec2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (i:l_iff a b), (forall (o:l_orec c a), l_orec c b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (i:l_iff a b) => (fun (o:l_orec c a) => l_oreci c b (l_iff_thor2 a b c i (l_orece1 c a o)) (l_iff_thec2 a b c i (l_orece2 c a o))))))). Time Defined. (* constant 112 *) Definition l_all : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (forall (x:sigma), p x))). Time Defined. (* constant 113 *) Definition l_alle : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (a1:l_all sigma p), (forall (s:sigma), p s)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (a1:l_all sigma p) => (fun (s:sigma) => a1 s)))). Time Defined. (* constant 114 *) Definition l_all_th1 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (n:l_not (p s)), l_not (l_all sigma p))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:sigma) => (fun (n:l_not (p s)) => (fun (x:l_all sigma p) => n (x s)))))). Time Defined. (* constant 115 *) Definition l_non : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (x:sigma), Prop))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (x:sigma) => l_not (p x)))). Time Defined. (* constant 116 *) Definition l_none : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => l_all sigma (l_non sigma p))). Time Defined. (* constant 117 *) Definition l_some : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => l_not (l_none sigma p))). Time Defined. (* constant 118 *) Definition l_somei : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (sp:p s), l_some sigma p)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:sigma) => (fun (sp:p s) => l_all_th1 sigma (l_non sigma p) s (l_weli (p s) sp))))). Time Defined. (* constant 119 *) Definition l_some_t1 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (n:l_not (l_all sigma p)), (forall (m:l_none sigma (l_non sigma p)), (forall (s:sigma), p s))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (n:l_not (l_all sigma p)) => (fun (m:l_none sigma (l_non sigma p)) => (fun (s:sigma) => l_et (p s) (m s)))))). Time Defined. (* constant 120 *) Definition l_some_t2 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (n:l_not (l_all sigma p)), (forall (m:l_none sigma (l_non sigma p)), l_con)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (n:l_not (l_all sigma p)) => (fun (m:l_none sigma (l_non sigma p)) => n (fun (x:sigma) => l_some_t1 sigma p n m x))))). Time Defined. (* constant 121 *) Definition l_some_th1 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (n:l_not (l_all sigma p)), l_some sigma (l_non sigma p)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (n:l_not (l_all sigma p)) => (fun (x:l_none sigma (l_non sigma p)) => l_some_t2 sigma p n x)))). Time Defined. (* constant 122 *) Definition l_some_t3 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:l_some sigma (l_non sigma p)), (forall (a1:l_all sigma p), (forall (t:sigma), l_not (l_not (p t))))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:l_some sigma (l_non sigma p)) => (fun (a1:l_all sigma p) => (fun (t:sigma) => l_weli (p t) (a1 t)))))). Time Defined. (* constant 123 *) Definition l_some_t4 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:l_some sigma (l_non sigma p)), (forall (a1:l_all sigma p), l_con)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:l_some sigma (l_non sigma p)) => (fun (a1:l_all sigma p) => s (fun (x:sigma) => l_some_t3 sigma p s a1 x))))). Time Defined. (* constant 124 *) Definition l_some_th2 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:l_some sigma (l_non sigma p)), l_not (l_all sigma p)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:l_some sigma (l_non sigma p)) => (fun (x:l_all sigma p) => l_some_t4 sigma p s x)))). Time Defined. (* constant 125 *) Definition l_some_th3 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (n:l_not (l_some sigma p)), l_none sigma p))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (n:l_not (l_some sigma p)) => l_et (l_none sigma p) n))). Time Defined. (* constant 126 *) Definition l_some_th4 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (n:l_not (l_some sigma p)), (forall (s:sigma), l_not (p s))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (n:l_not (l_some sigma p)) => (fun (s:sigma) => l_some_th3 sigma p n s)))). Time Defined. (* constant 127 *) Definition l_some_th5 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (n:l_none sigma p), l_not (l_some sigma p)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (n:l_none sigma p) => l_weli (l_none sigma p) n))). Time Defined. (* constant 128 *) Definition l_some_t5 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:l_some sigma p), (forall (x:Prop), (forall (i:(forall (y:sigma), l_imp (p y) x)), (forall (n:l_not x), (forall (t:sigma), l_not (p t)))))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:l_some sigma p) => (fun (x:Prop) => (fun (i:(forall (y:sigma), l_imp (p y) x)) => (fun (n:l_not x) => (fun (t:sigma) => l_imp_th3 (p t) x n (i t)))))))). Time Defined. (* constant 129 *) Definition l_some_t6 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:l_some sigma p), (forall (x:Prop), (forall (i:(forall (y:sigma), l_imp (p y) x)), (forall (n:l_not x), l_con)))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:l_some sigma p) => (fun (x:Prop) => (fun (i:(forall (y:sigma), l_imp (p y) x)) => (fun (n:l_not x) => l_mp (l_some sigma p) l_con s (l_some_th5 sigma p (fun (y:sigma) => l_some_t5 sigma p s x i n y)))))))). Time Defined. (* constant 130 *) Definition l_someapp : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:l_some sigma p), (forall (x:Prop), (forall (i:(forall (y:sigma), l_imp (p y) x)), x))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:l_some sigma p) => (fun (x:Prop) => (fun (i:(forall (y:sigma), l_imp (p y) x)) => l_et x (fun (y:l_not x) => l_some_t6 sigma p s x i y)))))). Time Defined. (* constant 131 *) Definition l_some_th6 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (q:(forall (x:sigma), Prop)), (forall (s:l_some sigma p), (forall (i:(forall (x:sigma), l_imp (p x) (q x))), l_some sigma q))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (q:(forall (x:sigma), Prop)) => (fun (s:l_some sigma p) => (fun (i:(forall (x:sigma), l_imp (p x) (q x))) => l_someapp sigma p s (l_some sigma q) (fun (x:sigma) => (fun (y:p x) => l_somei sigma q x (l_mp (p x) (q x) y (i x))))))))). Time Defined. (* constant 132 *) Definition l_or3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), Prop))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => l_or a (l_or b c)))). Time Defined. (* constant 133 *) Definition l_or3_th1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), (forall (n:l_not a), l_or b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => (fun (n:l_not a) => l_ore2 a (l_or b c) o n))))). Time Defined. (* constant 134 *) Definition l_or3e3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), (forall (n:l_not a), (forall (m:l_not b), c)))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => (fun (n:l_not a) => (fun (m:l_not b) => l_ore2 b c (l_or3_th1 a b c o n) m)))))). Time Defined. (* constant 135 *) Definition l_or3_th2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), (forall (n:l_not b), l_or c a))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => (fun (n:l_not b) => l_or_th2 c a (fun (x:l_not a) => l_or3e3 a b c o x n)))))). Time Defined. (* constant 136 *) Definition l_or3e1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), (forall (n:l_not b), (forall (m:l_not c), a)))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => (fun (n:l_not b) => (fun (m:l_not c) => l_ore2 c a (l_or3_th2 a b c o n) m)))))). Time Defined. (* constant 137 *) Definition l_or3_th3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), (forall (n:l_not c), l_or a b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => (fun (n:l_not c) => l_or_th2 a b (fun (x:l_not b) => l_or3e1 a b c o x n)))))). Time Defined. (* constant 138 *) Definition l_or3e2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), (forall (n:l_not c), (forall (m:l_not a), b)))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => (fun (n:l_not c) => (fun (m:l_not a) => l_ore2 a b (l_or3_th3 a b c o n) m)))))). Time Defined. (* constant 139 *) Definition l_or3_th4 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), l_or3 b c a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => l_or_th1 b (l_or c a) (fun (x:l_not b) => l_or3_th2 a b c o x))))). Time Defined. (* constant 140 *) Definition l_or3_th5 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), l_or3 c a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => l_or3_th4 b c a (l_or3_th4 a b c o))))). Time Defined. (* constant 141 *) Definition l_or3i1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:a), l_or3 a b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:a) => l_ori1 a (l_or b c) a1)))). Time Defined. (* constant 142 *) Definition l_or3i2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (b1:b), l_or3 a b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (b1:b) => l_ori2 a (l_or b c) (l_ori1 b c b1))))). Time Defined. (* constant 143 *) Definition l_or3i3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (c1:c), l_or3 a b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (c1:c) => l_ori2 a (l_or b c) (l_ori2 b c c1))))). Time Defined. (* constant 144 *) Definition l_or3_th6 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or a b), l_or3 a b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or a b) => l_or3_th4 c a b (l_ori2 c (l_or a b) o))))). Time Defined. (* constant 145 *) Definition l_or3_th7 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or b c), l_or3 a b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or b c) => l_ori2 a (l_or b c) o)))). Time Defined. (* constant 146 *) Definition l_or3_th8 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or c a), l_or3 a b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or c a) => l_or3_th4 c a b (l_or3_th6 c a b o))))). Time Defined. (* constant 147 *) Definition l_or3app : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (d:Prop), (forall (o:l_or3 a b c), (forall (i:l_imp a d), (forall (j:l_imp b d), (forall (k:l_imp c d), d)))))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (d:Prop) => (fun (o:l_or3 a b c) => (fun (i:l_imp a d) => (fun (j:l_imp b d) => (fun (k:l_imp c d) => l_orapp a (l_or b c) d o i (fun (x:l_or b c) => l_orapp b c d x j k))))))))). Time Defined. (* constant 148 *) Definition l_and3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), Prop))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => l_and a (l_and b c)))). Time Defined. (* constant 149 *) Definition l_and3e1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_ande1 a (l_and b c) a1)))). Time Defined. (* constant 150 *) Definition l_and3e2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_ande1 b c (l_ande2 a (l_and b c) a1))))). Time Defined. (* constant 151 *) Definition l_and3e3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_ande2 b c (l_ande2 a (l_and b c) a1))))). Time Defined. (* constant 152 *) Definition l_and3i : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:a), (forall (b1:b), (forall (c1:c), l_and3 a b c)))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:a) => (fun (b1:b) => (fun (c1:c) => l_andi a (l_and b c) a1 (l_andi b c b1 c1))))))). Time Defined. (* constant 153 *) Definition l_and3_th1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and3 b c a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_and3i b c a (l_and3e2 a b c a1) (l_and3e3 a b c a1) (l_and3e1 a b c a1))))). Time Defined. (* constant 154 *) Definition l_and3_th2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and3 c a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_and3_th1 b c a (l_and3_th1 a b c a1))))). Time Defined. (* constant 155 *) Definition l_and3_th3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_andi a b (l_and3e1 a b c a1) (l_and3e2 a b c a1))))). Time Defined. (* constant 156 *) Definition l_and3_th4 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_ande2 a (l_and b c) a1)))). Time Defined. (* constant 157 *) Definition l_and3_th5 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and c a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_and3_th3 c a b (l_and3_th2 a b c a1))))). Time Defined. (* constant 158 *) Definition l_and3_th6 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and3 c b a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:l_and3 a b c) => l_and3i c b a (l_and3e3 a b c a1) (l_and3e2 a b c a1) (l_and3e1 a b c a1))))). Time Defined. (* constant 159 *) Definition l_ec3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), Prop))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => l_and3 (l_ec a b) (l_ec b c) (l_ec c a)))). Time Defined. (* constant 160 *) Definition l_ec3_th1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => l_and3e1 (l_ec a b) (l_ec b c) (l_ec c a) e)))). Time Defined. (* constant 161 *) Definition l_ec3_th2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => l_and3e2 (l_ec a b) (l_ec b c) (l_ec c a) e)))). Time Defined. (* constant 162 *) Definition l_ec3_th3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec c a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => l_and3e3 (l_ec a b) (l_ec b c) (l_ec c a) e)))). Time Defined. (* constant 163 *) Definition l_ec3_th4 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec3 b c a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => l_and3_th1 (l_ec a b) (l_ec b c) (l_ec c a) e)))). Time Defined. (* constant 164 *) Definition l_ec3_th5 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec3 c a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => l_ec3_th4 b c a (l_ec3_th4 a b c e))))). Time Defined. (* constant 165 *) Definition l_ec3_th5a : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec3 c b a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => l_and3i (l_ec c b) (l_ec b a) (l_ec a c) (l_comec b c (l_ec3_th2 a b c e)) (l_comec a b (l_ec3_th1 a b c e)) (l_comec c a (l_ec3_th3 a b c e)))))). Time Defined. (* constant 166 *) Definition l_ec3e12 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (a1:a), l_not b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (a1:a) => l_ece1 a b (l_ec3_th1 a b c e) a1))))). Time Defined. (* constant 167 *) Definition l_ec3e13 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (a1:a), l_not c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (a1:a) => l_ece2 c a (l_ec3_th3 a b c e) a1))))). Time Defined. (* constant 168 *) Definition l_ec3e23 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (b1:b), l_not c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (b1:b) => l_ec3e12 b c a (l_ec3_th4 a b c e) b1))))). Time Defined. (* constant 169 *) Definition l_ec3e21 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (b1:b), l_not a))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (b1:b) => l_ec3e13 b c a (l_ec3_th4 a b c e) b1))))). Time Defined. (* constant 170 *) Definition l_ec3e31 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (c1:c), l_not a))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (c1:c) => l_ec3e12 c a b (l_ec3_th5 a b c e) c1))))). Time Defined. (* constant 171 *) Definition l_ec3e32 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (c1:c), l_not b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (c1:c) => l_ec3e13 c a b (l_ec3_th5 a b c e) c1))))). Time Defined. (* constant 172 *) Definition l_ec3_th6 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec a b), (forall (f:l_ec b c), (forall (g:l_ec c a), l_ec3 a b c)))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec a b) => (fun (f:l_ec b c) => (fun (g:l_ec c a) => l_and3i (l_ec a b) (l_ec b c) (l_ec c a) e f g)))))). Time Defined. (* constant 173 *) Definition l_ec3_th7 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (o:l_or a b), l_not c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (o:l_or a b) => l_orapp a b (l_not c) o (fun (x:a) => l_ece2 c a (l_ec3_th3 a b c e) x) (fun (x:b) => l_ece1 b c (l_ec3_th2 a b c e) x)))))). Time Defined. (* constant 174 *) Definition l_ec3_th8 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (o:l_or b c), l_not a))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (o:l_or b c) => l_ec3_th7 b c a (l_ec3_th4 a b c e) o))))). Time Defined. (* constant 175 *) Definition l_ec3_th9 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), (forall (o:l_or c a), l_not b))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (e:l_ec3 a b c) => (fun (o:l_or c a) => l_ec3_th7 c a b (l_ec3_th5 a b c e) o))))). Time Defined. (* constant 176 *) Definition l_ec3i1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (n:l_not a), (forall (m:l_not b), l_ec3 a b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (n:l_not a) => (fun (m:l_not b) => l_ec3_th6 a b c (l_eci1 a b n) (l_eci1 b c m) (l_eci2 c a n)))))). Time Defined. (* constant 177 *) Definition l_ec3i2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (n:l_not b), (forall (m:l_not c), l_ec3 a b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (n:l_not b) => (fun (m:l_not c) => l_ec3_th6 a b c (l_eci2 a b n) (l_eci1 b c n) (l_eci1 c a m)))))). Time Defined. (* constant 178 *) Definition l_ec3i3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (n:l_not c), (forall (m:l_not a), l_ec3 a b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (n:l_not c) => (fun (m:l_not a) => l_ec3_th6 a b c (l_eci1 a b m) (l_eci2 b c n) (l_eci1 c a n)))))). Time Defined. (* constant 179 *) Definition l_ec3_t1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (d:Prop), (forall (e:Prop), (forall (f:Prop), (forall (o1:l_or3 a b c), (forall (p1:l_ec3 d e f), (forall (i:l_imp a d), (forall (j:l_imp b e), (forall (k:l_imp c f), (forall (d1:d), l_not b)))))))))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (d:Prop) => (fun (e:Prop) => (fun (f:Prop) => (fun (o1:l_or3 a b c) => (fun (p1:l_ec3 d e f) => (fun (i:l_imp a d) => (fun (j:l_imp b e) => (fun (k:l_imp c f) => (fun (d1:d) => (fun (x:b) => l_mp e l_con (j x) (l_ec3e12 d e f p1 d1)))))))))))))). Time Defined. (* constant 180 *) Definition l_ec3_t2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (d:Prop), (forall (e:Prop), (forall (f:Prop), (forall (o1:l_or3 a b c), (forall (p1:l_ec3 d e f), (forall (i:l_imp a d), (forall (j:l_imp b e), (forall (k:l_imp c f), (forall (d1:d), l_not c)))))))))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (d:Prop) => (fun (e:Prop) => (fun (f:Prop) => (fun (o1:l_or3 a b c) => (fun (p1:l_ec3 d e f) => (fun (i:l_imp a d) => (fun (j:l_imp b e) => (fun (k:l_imp c f) => (fun (d1:d) => (fun (x:c) => l_mp f l_con (k x) (l_ec3e13 d e f p1 d1)))))))))))))). Time Defined. (* constant 181 *) Definition l_ec3_th10 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (d:Prop), (forall (e:Prop), (forall (f:Prop), (forall (o1:l_or3 a b c), (forall (p1:l_ec3 d e f), (forall (i:l_imp a d), (forall (j:l_imp b e), (forall (k:l_imp c f), (forall (d1:d), a)))))))))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (d:Prop) => (fun (e:Prop) => (fun (f:Prop) => (fun (o1:l_or3 a b c) => (fun (p1:l_ec3 d e f) => (fun (i:l_imp a d) => (fun (j:l_imp b e) => (fun (k:l_imp c f) => (fun (d1:d) => l_or3e1 a b c o1 (l_ec3_t1 a b c d e f o1 p1 i j k d1) (l_ec3_t2 a b c d e f o1 p1 i j k d1))))))))))))). Time Defined. (* constant 182 *) Definition l_ec3_th11 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (d:Prop), (forall (e:Prop), (forall (f:Prop), (forall (o1:l_or3 a b c), (forall (p1:l_ec3 d e f), (forall (i:l_imp a d), (forall (j:l_imp b e), (forall (k:l_imp c f), (forall (e1:e), b)))))))))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (d:Prop) => (fun (e:Prop) => (fun (f:Prop) => (fun (o1:l_or3 a b c) => (fun (p1:l_ec3 d e f) => (fun (i:l_imp a d) => (fun (j:l_imp b e) => (fun (k:l_imp c f) => (fun (e1:e) => l_ec3_th10 b c a e f d (l_or3_th4 a b c o1) (l_ec3_th4 d e f p1) j k i e1)))))))))))). Time Defined. (* constant 183 *) Definition l_ec3_th12 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (d:Prop), (forall (e:Prop), (forall (f:Prop), (forall (o1:l_or3 a b c), (forall (p1:l_ec3 d e f), (forall (i:l_imp a d), (forall (j:l_imp b e), (forall (k:l_imp c f), (forall (f1:f), c)))))))))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (d:Prop) => (fun (e:Prop) => (fun (f:Prop) => (fun (o1:l_or3 a b c) => (fun (p1:l_ec3 d e f) => (fun (i:l_imp a d) => (fun (j:l_imp b e) => (fun (k:l_imp c f) => (fun (f1:f) => l_ec3_th10 c a b f d e (l_or3_th5 a b c o1) (l_ec3_th5 d e f p1) k i j f1)))))))))))). Time Defined. (* constant 184 *) Definition l_orec3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), Prop))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => l_and (l_or3 a b c) (l_ec3 a b c)))). Time Defined. (* constant 185 *) Definition l_orec3e1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_orec3 a b c), l_or3 a b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_orec3 a b c) => l_ande1 (l_or3 a b c) (l_ec3 a b c) o)))). Time Defined. (* constant 186 *) Definition l_orec3e2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_orec3 a b c), l_ec3 a b c)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_orec3 a b c) => l_ande2 (l_or3 a b c) (l_ec3 a b c) o)))). Time Defined. (* constant 187 *) Definition l_orec3i : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or3 a b c), (forall (e:l_ec3 a b c), l_orec3 a b c))))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or3 a b c) => (fun (e:l_ec3 a b c) => l_andi (l_or3 a b c) (l_ec3 a b c) o e))))). Time Defined. (* constant 188 *) Definition l_orec3_th1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_orec3 a b c), l_orec3 b c a)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_orec3 a b c) => l_orec3i b c a (l_or3_th4 a b c (l_orec3e1 a b c o)) (l_ec3_th4 a b c (l_orec3e2 a b c o)))))). Time Defined. (* constant 189 *) Definition l_orec3_th2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_orec3 a b c), l_orec3 c a b)))). exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_orec3 a b c) => l_orec3i c a b (l_or3_th5 a b c (l_orec3e1 a b c o)) (l_ec3_th5 a b c (l_orec3e2 a b c o)))))). Time Defined. (* constant 190 *) Axiom l_e_is : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), Prop))). (* constant 191 *) Axiom l_e_refis : (forall (sigma:Type), (forall (s:sigma), l_e_is sigma s s)). (* constant 192 *) Axiom l_e_isp : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (t:sigma), (forall (sp:p s), (forall (i:l_e_is sigma s t), p t)))))). (* constant 193 *) Definition l_e_symis : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma s t), l_e_is sigma t s)))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma s t) => l_e_isp sigma (fun (x:sigma) => l_e_is sigma x s) s t (l_e_refis sigma s) i)))). Time Defined. (* constant 194 *) Definition l_e_tris : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (i:l_e_is sigma s t), (forall (j:l_e_is sigma t u), l_e_is sigma s u)))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (i:l_e_is sigma s t) => (fun (j:l_e_is sigma t u) => l_e_isp sigma (fun (x:sigma) => l_e_is sigma x u) t s j (l_e_symis sigma s t i))))))). Time Defined. (* constant 195 *) Definition l_e_tris1 : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (i:l_e_is sigma u s), (forall (j:l_e_is sigma u t), l_e_is sigma s t)))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (i:l_e_is sigma u s) => (fun (j:l_e_is sigma u t) => l_e_tris sigma s u t (l_e_symis sigma u s i) j)))))). Time Defined. (* constant 196 *) Definition l_e_tris2 : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (i:l_e_is sigma s u), (forall (j:l_e_is sigma t u), l_e_is sigma s t)))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (i:l_e_is sigma s u) => (fun (j:l_e_is sigma t u) => l_e_tris sigma s u t i (l_e_symis sigma t u j))))))). Time Defined. (* constant 197 *) Definition l_e_isp1 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (t:sigma), (forall (sp:p s), (forall (i:l_e_is sigma t s), p t)))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:sigma) => (fun (t:sigma) => (fun (sp:p s) => (fun (i:l_e_is sigma t s) => l_e_isp sigma p s t sp (l_e_symis sigma t s i))))))). Time Defined. (* constant 198 *) Definition l_e_symnotis : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (n:l_not (l_e_is sigma s t)), l_not (l_e_is sigma t s))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (n:l_not (l_e_is sigma s t)) => l_imp_th3 (l_e_is sigma t s) (l_e_is sigma s t) n (fun (x:l_e_is sigma t s) => l_e_symis sigma t s x))))). Time Defined. (* constant 199 *) Definition l_e_notis_th1 : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (n:l_not (l_e_is sigma s t)), (forall (i:l_e_is sigma s u), l_not (l_e_is sigma u t))))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (n:l_not (l_e_is sigma s t)) => (fun (i:l_e_is sigma s u) => l_e_isp sigma (fun (x:sigma) => l_not (l_e_is sigma x t)) s u n i)))))). Time Defined. (* constant 200 *) Definition l_e_notis_th2 : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (n:l_not (l_e_is sigma s t)), (forall (i:l_e_is sigma u s), l_not (l_e_is sigma u t))))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (n:l_not (l_e_is sigma s t)) => (fun (i:l_e_is sigma u s) => l_e_notis_th1 sigma s t u n (l_e_symis sigma u s i))))))). Time Defined. (* constant 201 *) Definition l_e_notis_th3 : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (n:l_not (l_e_is sigma s t)), (forall (i:l_e_is sigma t u), l_not (l_e_is sigma s u))))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (n:l_not (l_e_is sigma s t)) => (fun (i:l_e_is sigma t u) => l_e_isp sigma (fun (x:sigma) => l_not (l_e_is sigma s x)) t u n i)))))). Time Defined. (* constant 202 *) Definition l_e_notis_th4 : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (n:l_not (l_e_is sigma s t)), (forall (i:l_e_is sigma u t), l_not (l_e_is sigma s u))))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (n:l_not (l_e_is sigma s t)) => (fun (i:l_e_is sigma u t) => l_e_notis_th3 sigma s t u n (l_e_symis sigma u t i))))))). Time Defined. (* constant 203 *) Definition l_e_notis_th5 : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (v:sigma), (forall (n:l_not (l_e_is sigma s t)), (forall (i:l_e_is sigma s u), (forall (j:l_e_is sigma t v), l_not (l_e_is sigma u v))))))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (v:sigma) => (fun (n:l_not (l_e_is sigma s t)) => (fun (i:l_e_is sigma s u) => (fun (j:l_e_is sigma t v) => l_e_notis_th3 sigma u t v (l_e_notis_th1 sigma s t u n i) j)))))))). Time Defined. (* constant 204 *) Definition l_e_tr3is : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (v:sigma), (forall (i:l_e_is sigma s t), (forall (j:l_e_is sigma t u), (forall (k:l_e_is sigma u v), l_e_is sigma s v)))))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (v:sigma) => (fun (i:l_e_is sigma s t) => (fun (j:l_e_is sigma t u) => (fun (k:l_e_is sigma u v) => l_e_tris sigma s u v (l_e_tris sigma s t u i j) k)))))))). Time Defined. (* constant 205 *) Definition l_e_tr4is : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (v:sigma), (forall (w:sigma), (forall (i:l_e_is sigma s t), (forall (j:l_e_is sigma t u), (forall (k:l_e_is sigma u v), (forall (l:l_e_is sigma v w), l_e_is sigma s w)))))))))). exact (fun (sigma:Type) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (v:sigma) => (fun (w:sigma) => (fun (i:l_e_is sigma s t) => (fun (j:l_e_is sigma t u) => (fun (k:l_e_is sigma u v) => (fun (l:l_e_is sigma v w) => l_e_tris sigma s v w (l_e_tr3is sigma s t u v i j k) l)))))))))). Time Defined. (* constant 206 *) Definition l_e_amone : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (forall (x:sigma), (forall (y:sigma), (forall (u:p x), (forall (v:p y), l_e_is sigma x y)))))). Time Defined. (* constant 207 *) Definition l_e_amonee : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (a1:l_e_amone sigma p), (forall (s:sigma), (forall (t:sigma), (forall (sp:p s), (forall (tp:p t), l_e_is sigma s t))))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (a1:l_e_amone sigma p) => (fun (s:sigma) => (fun (t:sigma) => (fun (sp:p s) => (fun (tp:p t) => a1 s t sp tp))))))). Time Defined. (* constant 208 *) Definition l_e_one : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => l_and (l_e_amone sigma p) (l_some sigma p))). Time Defined. (* constant 209 *) Definition l_e_onei : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (a1:l_e_amone sigma p), (forall (s:l_some sigma p), l_e_one sigma p)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (a1:l_e_amone sigma p) => (fun (s:l_some sigma p) => l_andi (l_e_amone sigma p) (l_some sigma p) a1 s)))). Time Defined. (* constant 210 *) Definition l_e_onee1 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_one sigma p), l_e_amone sigma p))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (o1:l_e_one sigma p) => l_ande1 (l_e_amone sigma p) (l_some sigma p) o1))). Time Defined. (* constant 211 *) Definition l_e_onee2 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_one sigma p), l_some sigma p))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (o1:l_e_one sigma p) => l_ande2 (l_e_amone sigma p) (l_some sigma p) o1))). Time Defined. (* constant 212 *) Axiom l_e_ind : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_one sigma p), sigma))). (* constant 213 *) Axiom l_e_oneax : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_one sigma p), p (l_e_ind sigma p o1)))). (* constant 214 *) Definition l_e_one_th1 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_one sigma p), (forall (s:sigma), (forall (sp:p s), l_e_is sigma (l_e_ind sigma p o1) s))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (o1:l_e_one sigma p) => (fun (s:sigma) => (fun (sp:p s) => l_e_amonee sigma p (l_e_onee1 sigma p o1) (l_e_ind sigma p o1) s (l_e_oneax sigma p o1) sp))))). Time Defined. (* constant 215 *) Definition l_e_isf : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma s t), l_e_is tau (f s) (f t))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma s t) => l_e_isp sigma (fun (x:sigma) => l_e_is tau (f s) (f x)) s t (l_e_refis tau (f s)) i)))))). Time Defined. (* constant 216 *) Definition l_e_injective : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), Prop))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => l_all sigma (fun (x:sigma) => l_all sigma (fun (y:sigma) => l_imp (l_e_is tau (f x) (f y)) (l_e_is sigma x y)))))). Time Defined. (* constant 217 *) Definition l_e_isfe : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (s:sigma), (forall (t:sigma), (forall (j:l_e_is tau (f s) (f t)), l_e_is sigma s t))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (s:sigma) => (fun (t:sigma) => (fun (j:l_e_is tau (f s) (f t)) => i s t j))))))). Time Defined. (* constant 218 *) Definition l_e_image : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (t0:tau), Prop)))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (t0:tau) => l_some sigma (fun (x:sigma) => l_e_is tau t0 (f x)))))). Time Defined. (* constant 219 *) Definition l_e_tofs : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s:sigma), tau)))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s:sigma) => f s)))). Time Defined. (* constant 220 *) Definition l_e_imagei : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s:sigma), l_e_image sigma tau f (l_e_tofs sigma tau f s))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s:sigma) => l_somei sigma (fun (x:sigma) => l_e_is tau (l_e_tofs sigma tau f s) (f x)) s (l_e_refis tau (l_e_tofs sigma tau f s)))))). Time Defined. (* constant 221 *) Definition l_e_inj_t1 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (ta:tau), (forall (tb:tau), (forall (j:l_e_is tau ta tb), (forall (sa:sigma), (forall (sb:sigma), (forall (ja:l_e_is tau ta (l_e_tofs sigma tau f sa)), (forall (jb:l_e_is tau tb (l_e_tofs sigma tau f sb)), l_e_is tau (l_e_tofs sigma tau f sa) (l_e_tofs sigma tau f sb)))))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (ta:tau) => (fun (tb:tau) => (fun (j:l_e_is tau ta tb) => (fun (sa:sigma) => (fun (sb:sigma) => (fun (ja:l_e_is tau ta (l_e_tofs sigma tau f sa)) => (fun (jb:l_e_is tau tb (l_e_tofs sigma tau f sb)) => l_e_tr3is tau (l_e_tofs sigma tau f sa) ta tb (l_e_tofs sigma tau f sb) (l_e_symis tau ta (l_e_tofs sigma tau f sa) ja) j jb))))))))))). Time Defined. (* constant 222 *) Definition l_e_inj_th1 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (ta:tau), (forall (tb:tau), (forall (j:l_e_is tau ta tb), (forall (sa:sigma), (forall (sb:sigma), (forall (ja:l_e_is tau ta (l_e_tofs sigma tau f sa)), (forall (jb:l_e_is tau tb (l_e_tofs sigma tau f sb)), l_e_is sigma sa sb))))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (ta:tau) => (fun (tb:tau) => (fun (j:l_e_is tau ta tb) => (fun (sa:sigma) => (fun (sb:sigma) => (fun (ja:l_e_is tau ta (l_e_tofs sigma tau f sa)) => (fun (jb:l_e_is tau tb (l_e_tofs sigma tau f sb)) => l_e_isfe sigma tau f i sa sb (l_e_inj_t1 sigma tau f i ta tb j sa sb ja jb)))))))))))). Time Defined. (* constant 223 *) Definition l_e_inj_th2 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (t0:tau), l_e_amone sigma (fun (x:sigma) => l_e_is tau t0 (f x))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (t0:tau) => (fun (x:sigma) => (fun (y:sigma) => (fun (u:l_e_is tau t0 (f x)) => (fun (v:l_e_is tau t0 (f y)) => l_e_inj_th1 sigma tau f i t0 t0 (l_e_refis tau t0) x y u v))))))))). Time Defined. (* constant 224 *) Definition l_e_inj_th3 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (t0:tau), (forall (j:l_e_image sigma tau f t0), l_e_one sigma (fun (x:sigma) => l_e_is tau t0 (f x)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (t0:tau) => (fun (j:l_e_image sigma tau f t0) => l_e_onei sigma (fun (x:sigma) => l_e_is tau t0 (f x)) (l_e_inj_th2 sigma tau f i t0) j)))))). Time Defined. (* constant 225 *) Definition l_e_soft : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (t0:tau), (forall (j:l_e_image sigma tau f t0), sigma)))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (t0:tau) => (fun (j:l_e_image sigma tau f t0) => l_e_ind sigma (fun (x:sigma) => l_e_is tau t0 (f x)) (l_e_inj_th3 sigma tau f i t0 j))))))). Time Defined. (* constant 226 *) Definition l_e_inverse : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (x:tau), (forall (u:l_e_image sigma tau f x), sigma)))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (x:tau) => (fun (u:l_e_image sigma tau f x) => l_e_soft sigma tau f i x u)))))). Time Defined. (* constant 227 *) Definition l_e_ists1 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (t0:tau), (forall (j:l_e_image sigma tau f t0), l_e_is tau t0 (l_e_tofs sigma tau f (l_e_soft sigma tau f i t0 j)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (t0:tau) => (fun (j:l_e_image sigma tau f t0) => l_e_oneax sigma (fun (x:sigma) => l_e_is tau t0 (f x)) (l_e_inj_th3 sigma tau f i t0 j))))))). Time Defined. (* constant 228 *) Definition l_e_ists2 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (t0:tau), (forall (j:l_e_image sigma tau f t0), l_e_is tau (l_e_tofs sigma tau f (l_e_soft sigma tau f i t0 j)) t0)))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (t0:tau) => (fun (j:l_e_image sigma tau f t0) => l_e_symis tau t0 (l_e_tofs sigma tau f (l_e_soft sigma tau f i t0 j)) (l_e_ists1 sigma tau f i t0 j))))))). Time Defined. (* constant 229 *) Definition l_e_isinv : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (ta:tau), (forall (ja:l_e_image sigma tau f ta), (forall (tb:tau), (forall (jb:l_e_image sigma tau f tb), (forall (j:l_e_is tau ta tb), l_e_is sigma (l_e_soft sigma tau f i ta ja) (l_e_soft sigma tau f i tb jb)))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (ta:tau) => (fun (ja:l_e_image sigma tau f ta) => (fun (tb:tau) => (fun (jb:l_e_image sigma tau f tb) => (fun (j:l_e_is tau ta tb) => l_e_inj_th1 sigma tau f i ta tb j (l_e_soft sigma tau f i ta ja) (l_e_soft sigma tau f i tb jb) (l_e_ists1 sigma tau f i ta ja) (l_e_ists1 sigma tau f i tb jb)))))))))). Time Defined. (* constant 230 *) Definition l_e_isinve : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (ta:tau), (forall (ja:l_e_image sigma tau f ta), (forall (tb:tau), (forall (jb:l_e_image sigma tau f tb), (forall (j:l_e_is sigma (l_e_soft sigma tau f i ta ja) (l_e_soft sigma tau f i tb jb)), l_e_is tau ta tb))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (ta:tau) => (fun (ja:l_e_image sigma tau f ta) => (fun (tb:tau) => (fun (jb:l_e_image sigma tau f tb) => (fun (j:l_e_is sigma (l_e_soft sigma tau f i ta ja) (l_e_soft sigma tau f i tb jb)) => l_e_tr3is tau ta (l_e_tofs sigma tau f (l_e_soft sigma tau f i ta ja)) (l_e_tofs sigma tau f (l_e_soft sigma tau f i tb jb)) tb (l_e_ists1 sigma tau f i ta ja) (l_e_isf sigma tau f (l_e_soft sigma tau f i ta ja) (l_e_soft sigma tau f i tb jb) j) (l_e_ists2 sigma tau f i tb jb)))))))))). Time Defined. (* constant 231 *) Definition l_e_isst1 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (s:sigma), l_e_is sigma s (l_e_soft sigma tau f i (l_e_tofs sigma tau f s) (l_e_imagei sigma tau f s))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (s:sigma) => l_e_isfe sigma tau f i s (l_e_soft sigma tau f i (l_e_tofs sigma tau f s) (l_e_imagei sigma tau f s)) (l_e_ists1 sigma tau f i (l_e_tofs sigma tau f s) (l_e_imagei sigma tau f s))))))). Time Defined. (* constant 232 *) Definition l_e_isst2 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (i:l_e_injective sigma tau f), (forall (s:sigma), l_e_is sigma (l_e_soft sigma tau f i (l_e_tofs sigma tau f s) (l_e_imagei sigma tau f s)) s))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (i:l_e_injective sigma tau f) => (fun (s:sigma) => l_e_symis sigma s (l_e_soft sigma tau f i (l_e_tofs sigma tau f s) (l_e_imagei sigma tau f s)) (l_e_isst1 sigma tau f i s)))))). Time Defined. (* constant 233 *) Definition l_e_surjective : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), Prop))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => l_all tau (fun (x:tau) => l_e_image sigma tau f x)))). Time Defined. (* constant 234 *) Definition l_e_bijective : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), Prop))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => l_and (l_e_injective sigma tau f) (l_e_surjective sigma tau f)))). Time Defined. (* constant 235 *) Definition l_e_inj_t2 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (b:l_e_bijective sigma tau f), l_e_injective sigma tau f)))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (b:l_e_bijective sigma tau f) => l_ande1 (l_e_injective sigma tau f) (l_e_surjective sigma tau f) b)))). Time Defined. (* constant 236 *) Definition l_e_inj_t3 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (b:l_e_bijective sigma tau f), l_e_surjective sigma tau f)))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (b:l_e_bijective sigma tau f) => l_ande2 (l_e_injective sigma tau f) (l_e_surjective sigma tau f) b)))). Time Defined. (* constant 237 *) Definition l_e_inj_so : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (b:l_e_bijective sigma tau f), (forall (t:tau), sigma))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (b:l_e_bijective sigma tau f) => (fun (t:tau) => l_e_soft sigma tau f (l_e_inj_t2 sigma tau f b) t (l_e_inj_t3 sigma tau f b t)))))). Time Defined. (* constant 238 *) Definition l_e_invf : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (b:l_e_bijective sigma tau f), (forall (x:tau), sigma))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (b:l_e_bijective sigma tau f) => (fun (x:tau) => l_e_inj_so sigma tau f b x))))). Time Defined. (* constant 239 *) Definition l_e_thinvf1 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (b:l_e_bijective sigma tau f), (forall (s:sigma), l_e_is sigma s (l_e_invf sigma tau f b (f s))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (b:l_e_bijective sigma tau f) => (fun (s:sigma) => l_e_tris sigma s (l_e_soft sigma tau f (l_e_inj_t2 sigma tau f b) (l_e_tofs sigma tau f s) (l_e_imagei sigma tau f s)) (l_e_invf sigma tau f b (f s)) (l_e_isst1 sigma tau f (l_e_inj_t2 sigma tau f b) s) (l_e_isinv sigma tau f (l_e_inj_t2 sigma tau f b) (l_e_tofs sigma tau f s) (l_e_imagei sigma tau f s) (l_e_tofs sigma tau f s) (l_e_inj_t3 sigma tau f b (l_e_tofs sigma tau f s)) (l_e_refis tau (l_e_tofs sigma tau f s)))))))). Time Defined. (* constant 240 *) Definition l_e_thinvf2 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (b:l_e_bijective sigma tau f), (forall (t:tau), l_e_is tau t (f (l_e_invf sigma tau f b t))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (b:l_e_bijective sigma tau f) => (fun (t:tau) => l_e_ists1 sigma tau f (l_e_inj_t2 sigma tau f b) t (l_e_inj_t3 sigma tau f b t)))))). Time Defined. (* constant 241 *) Definition l_e_inj_h : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (nif:l_e_injective sigma tau f), (forall (ig:l_e_injective tau upsilon g), (forall (x:sigma), upsilon)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (nif:l_e_injective sigma tau f) => (fun (ig:l_e_injective tau upsilon g) => (fun (x:sigma) => g (f x))))))))). Time Defined. (* constant 242 *) Definition l_e_inj_t4 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (nif:l_e_injective sigma tau f), (forall (ig:l_e_injective tau upsilon g), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is upsilon (l_e_inj_h sigma tau upsilon f g nif ig s) (l_e_inj_h sigma tau upsilon f g nif ig t)), l_e_is tau (f s) (f t))))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (nif:l_e_injective sigma tau f) => (fun (ig:l_e_injective tau upsilon g) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is upsilon (l_e_inj_h sigma tau upsilon f g nif ig s) (l_e_inj_h sigma tau upsilon f g nif ig t)) => l_e_isfe tau upsilon g ig (f s) (f t) i)))))))))). Time Defined. (* constant 243 *) Definition l_e_inj_t5 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (nif:l_e_injective sigma tau f), (forall (ig:l_e_injective tau upsilon g), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is upsilon (l_e_inj_h sigma tau upsilon f g nif ig s) (l_e_inj_h sigma tau upsilon f g nif ig t)), l_e_is sigma s t)))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (nif:l_e_injective sigma tau f) => (fun (ig:l_e_injective tau upsilon g) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is upsilon (l_e_inj_h sigma tau upsilon f g nif ig s) (l_e_inj_h sigma tau upsilon f g nif ig t)) => l_e_isfe sigma tau f nif s t (l_e_inj_t4 sigma tau upsilon f g nif ig s t i))))))))))). Time Defined. (* constant 244 *) Definition l_e_inj_th4 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (nif:l_e_injective sigma tau f), (forall (ig:l_e_injective tau upsilon g), l_e_injective sigma upsilon (fun (x:sigma) => g (f x))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (nif:l_e_injective sigma tau f) => (fun (ig:l_e_injective tau upsilon g) => (fun (x:sigma) => (fun (y:sigma) => (fun (v:l_e_is upsilon (l_e_inj_h sigma tau upsilon f g nif ig x) (l_e_inj_h sigma tau upsilon f g nif ig y)) => l_e_inj_t5 sigma tau upsilon f g nif ig x y v)))))))))). Time Defined. (* constant 245 *) Definition l_e_surj_h : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (sf:l_e_surjective sigma tau f), (forall (sg:l_e_surjective tau upsilon g), (forall (x:sigma), upsilon)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (sf:l_e_surjective sigma tau f) => (fun (sg:l_e_surjective tau upsilon g) => (fun (x:sigma) => g (f x))))))))). Time Defined. (* constant 246 *) Definition l_e_surj_t1 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (sf:l_e_surjective sigma tau f), (forall (sg:l_e_surjective tau upsilon g), (forall (u:upsilon), l_e_image tau upsilon g u)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (sf:l_e_surjective sigma tau f) => (fun (sg:l_e_surjective tau upsilon g) => (fun (u:upsilon) => sg u)))))))). Time Defined. (* constant 247 *) Definition l_e_surj_t2 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (sf:l_e_surjective sigma tau f), (forall (sg:l_e_surjective tau upsilon g), (forall (u:upsilon), (forall (t:tau), (forall (i:l_e_is upsilon u (g t)), l_e_image sigma tau f t)))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (sf:l_e_surjective sigma tau f) => (fun (sg:l_e_surjective tau upsilon g) => (fun (u:upsilon) => (fun (t:tau) => (fun (i:l_e_is upsilon u (g t)) => sf t)))))))))). Time Defined. (* constant 248 *) Definition l_e_surj_t3 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (sf:l_e_surjective sigma tau f), (forall (sg:l_e_surjective tau upsilon g), (forall (u:upsilon), (forall (t:tau), (forall (i:l_e_is upsilon u (g t)), (forall (s:sigma), (forall (j:l_e_is tau t (f s)), l_e_is upsilon u (l_e_surj_h sigma tau upsilon f g sf sg s))))))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (sf:l_e_surjective sigma tau f) => (fun (sg:l_e_surjective tau upsilon g) => (fun (u:upsilon) => (fun (t:tau) => (fun (i:l_e_is upsilon u (g t)) => (fun (s:sigma) => (fun (j:l_e_is tau t (f s)) => l_e_tris upsilon u (g t) (l_e_surj_h sigma tau upsilon f g sf sg s) i (l_e_isf tau upsilon g t (f s) j))))))))))))). Time Defined. (* constant 249 *) Definition l_e_surj_t4 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (sf:l_e_surjective sigma tau f), (forall (sg:l_e_surjective tau upsilon g), (forall (u:upsilon), (forall (t:tau), (forall (i:l_e_is upsilon u (g t)), (forall (s:sigma), (forall (j:l_e_is tau t (f s)), l_e_image sigma upsilon (l_e_surj_h sigma tau upsilon f g sf sg) u)))))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (sf:l_e_surjective sigma tau f) => (fun (sg:l_e_surjective tau upsilon g) => (fun (u:upsilon) => (fun (t:tau) => (fun (i:l_e_is upsilon u (g t)) => (fun (s:sigma) => (fun (j:l_e_is tau t (f s)) => l_somei sigma (fun (x:sigma) => l_e_is upsilon u (l_e_surj_h sigma tau upsilon f g sf sg x)) s (l_e_surj_t3 sigma tau upsilon f g sf sg u t i s j))))))))))))). Time Defined. (* constant 250 *) Definition l_e_surj_t5 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (sf:l_e_surjective sigma tau f), (forall (sg:l_e_surjective tau upsilon g), (forall (u:upsilon), (forall (t:tau), (forall (i:l_e_is upsilon u (g t)), l_e_image sigma upsilon (l_e_surj_h sigma tau upsilon f g sf sg) u)))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (sf:l_e_surjective sigma tau f) => (fun (sg:l_e_surjective tau upsilon g) => (fun (u:upsilon) => (fun (t:tau) => (fun (i:l_e_is upsilon u (g t)) => l_someapp sigma (fun (x:sigma) => l_e_is tau t (f x)) (l_e_surj_t2 sigma tau upsilon f g sf sg u t i) (l_e_image sigma upsilon (l_e_surj_h sigma tau upsilon f g sf sg) u) (fun (x:sigma) => (fun (v:l_e_is tau t (f x)) => l_e_surj_t4 sigma tau upsilon f g sf sg u t i x v)))))))))))). Time Defined. (* constant 251 *) Definition l_e_surj_t6 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (sf:l_e_surjective sigma tau f), (forall (sg:l_e_surjective tau upsilon g), (forall (u:upsilon), l_e_image sigma upsilon (l_e_surj_h sigma tau upsilon f g sf sg) u)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (sf:l_e_surjective sigma tau f) => (fun (sg:l_e_surjective tau upsilon g) => (fun (u:upsilon) => l_someapp tau (fun (x:tau) => l_e_is upsilon u (g x)) (l_e_surj_t1 sigma tau upsilon f g sf sg u) (l_e_image sigma upsilon (l_e_surj_h sigma tau upsilon f g sf sg) u) (fun (x:tau) => (fun (v:l_e_is upsilon u (g x)) => l_e_surj_t5 sigma tau upsilon f g sf sg u x v)))))))))). Time Defined. (* constant 252 *) Definition l_e_surj_th1 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (sf:l_e_surjective sigma tau f), (forall (sg:l_e_surjective tau upsilon g), l_e_surjective sigma upsilon (fun (x:sigma) => g (f x))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (sf:l_e_surjective sigma tau f) => (fun (sg:l_e_surjective tau upsilon g) => (fun (x:upsilon) => l_e_surj_t6 sigma tau upsilon f g sf sg x)))))))). Time Defined. (* constant 253 *) Definition l_e_bij_h : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (bf:l_e_bijective sigma tau f), (forall (bg:l_e_bijective tau upsilon g), (forall (x:sigma), upsilon)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (bf:l_e_bijective sigma tau f) => (fun (bg:l_e_bijective tau upsilon g) => (fun (x:sigma) => g (f x))))))))). Time Defined. (* constant 254 *) Definition l_e_bij_t1 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (bf:l_e_bijective sigma tau f), (forall (bg:l_e_bijective tau upsilon g), l_e_injective sigma upsilon (l_e_bij_h sigma tau upsilon f g bf bg)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (bf:l_e_bijective sigma tau f) => (fun (bg:l_e_bijective tau upsilon g) => l_e_inj_th4 sigma tau upsilon f g (l_ande1 (l_e_injective sigma tau f) (l_e_surjective sigma tau f) bf) (l_ande1 (l_e_injective tau upsilon g) (l_e_surjective tau upsilon g) bg)))))))). Time Defined. (* constant 255 *) Definition l_e_bij_t2 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (bf:l_e_bijective sigma tau f), (forall (bg:l_e_bijective tau upsilon g), l_e_surjective sigma upsilon (l_e_bij_h sigma tau upsilon f g bf bg)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (bf:l_e_bijective sigma tau f) => (fun (bg:l_e_bijective tau upsilon g) => l_e_surj_th1 sigma tau upsilon f g (l_ande2 (l_e_injective sigma tau f) (l_e_surjective sigma tau f) bf) (l_ande2 (l_e_injective tau upsilon g) (l_e_surjective tau upsilon g) bg)))))))). Time Defined. (* constant 256 *) Definition l_e_bij_th1 : (forall (sigma:Type), (forall (tau:Type), (forall (upsilon:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:tau), upsilon)), (forall (bf:l_e_bijective sigma tau f), (forall (bg:l_e_bijective tau upsilon g), l_e_bijective sigma upsilon (fun (x:sigma) => g (f x))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (upsilon:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:tau), upsilon)) => (fun (bf:l_e_bijective sigma tau f) => (fun (bg:l_e_bijective tau upsilon g) => l_andi (l_e_injective sigma upsilon (l_e_bij_h sigma tau upsilon f g bf bg)) (l_e_surjective sigma upsilon (l_e_bij_h sigma tau upsilon f g bf bg)) (l_e_bij_t1 sigma tau upsilon f g bf bg) (l_e_bij_t2 sigma tau upsilon f g bf bg)))))))). Time Defined. (* constant 257 *) Definition l_e_fise : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:sigma), tau)), (forall (i:l_e_is (forall (x:sigma), tau) f g), (forall (s:sigma), l_e_is tau (f s) (g s))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:sigma), tau)) => (fun (i:l_e_is (forall (x:sigma), tau) f g) => (fun (s:sigma) => l_e_isp (forall (x:sigma), tau) (fun (y:(forall (x:sigma), tau)) => l_e_is tau (f s) (y s)) f g (l_e_refis tau (f s)) i)))))). Time Defined. (* constant 258 *) Axiom l_e_fisi : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:sigma), tau)), (forall (i:(forall (x:sigma), l_e_is tau (f x) (g x))), l_e_is (forall (x:sigma), tau) f g))))). (* constant 259 *) Definition l_e_fis_th1 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (g:(forall (x:sigma), tau)), (forall (i:l_e_is (forall (x:sigma), tau) f g), (forall (s:sigma), (forall (t:sigma), (forall (j:l_e_is sigma s t), l_e_is tau (f s) (g t))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (g:(forall (x:sigma), tau)) => (fun (i:l_e_is (forall (x:sigma), tau) f g) => (fun (s:sigma) => (fun (t:sigma) => (fun (j:l_e_is sigma s t) => l_e_tris tau (f s) (f t) (g t) (l_e_isf sigma tau f s t j) (l_e_fise sigma tau f g i t))))))))). Time Defined. (* constant 260 *) Axiom l_e_ot : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Type)). (* constant 261 *) Axiom l_e_in : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_ot sigma p), sigma))). (* constant 262 *) Axiom l_e_inp : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_ot sigma p), p (l_e_in sigma p o1)))). (* constant 263 *) Axiom l_e_otax1 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), l_e_injective (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x))). (* constant 264 *) Axiom l_e_otax2 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (sp:p s), l_e_image (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) s)))). (* constant 265 *) Definition l_e_isini : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_ot sigma p), (forall (o2:l_e_ot sigma p), (forall (i:l_e_is (l_e_ot sigma p) o1 o2), l_e_is sigma (l_e_in sigma p o1) (l_e_in sigma p o2)))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (o1:l_e_ot sigma p) => (fun (o2:l_e_ot sigma p) => (fun (i:l_e_is (l_e_ot sigma p) o1 o2) => l_e_isf (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) o1 o2 i))))). Time Defined. (* constant 266 *) Definition l_e_isine : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_ot sigma p), (forall (o2:l_e_ot sigma p), (forall (i:l_e_is sigma (l_e_in sigma p o1) (l_e_in sigma p o2)), l_e_is (l_e_ot sigma p) o1 o2))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (o1:l_e_ot sigma p) => (fun (o2:l_e_ot sigma p) => (fun (i:l_e_is sigma (l_e_in sigma p o1) (l_e_in sigma p o2)) => l_e_isfe (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) (l_e_otax1 sigma p) o1 o2 i))))). Time Defined. (* constant 267 *) Definition l_e_out : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (sp:p s), l_e_ot sigma p)))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:sigma) => (fun (sp:p s) => l_e_soft (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) (l_e_otax1 sigma p) s (l_e_otax2 sigma p s sp))))). Time Defined. (* constant 268 *) Definition l_e_isouti : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (sp:p s), (forall (t:sigma), (forall (tp:p t), (forall (i:l_e_is sigma s t), l_e_is (l_e_ot sigma p) (l_e_out sigma p s sp) (l_e_out sigma p t tp)))))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:sigma) => (fun (sp:p s) => (fun (t:sigma) => (fun (tp:p t) => (fun (i:l_e_is sigma s t) => l_e_isinv (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) (l_e_otax1 sigma p) s (l_e_otax2 sigma p s sp) t (l_e_otax2 sigma p t tp) i))))))). Time Defined. (* constant 269 *) Definition l_e_isoute : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (sp:p s), (forall (t:sigma), (forall (tp:p t), (forall (i:l_e_is (l_e_ot sigma p) (l_e_out sigma p s sp) (l_e_out sigma p t tp)), l_e_is sigma s t))))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:sigma) => (fun (sp:p s) => (fun (t:sigma) => (fun (tp:p t) => (fun (i:l_e_is (l_e_ot sigma p) (l_e_out sigma p s sp) (l_e_out sigma p t tp)) => l_e_isinve (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) (l_e_otax1 sigma p) s (l_e_otax2 sigma p s sp) t (l_e_otax2 sigma p t tp) i))))))). Time Defined. (* constant 270 *) Definition l_e_isoutin : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_ot sigma p), l_e_is (l_e_ot sigma p) o1 (l_e_out sigma p (l_e_in sigma p o1) (l_e_inp sigma p o1))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (o1:l_e_ot sigma p) => l_e_tris (l_e_ot sigma p) o1 (l_e_soft (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) (l_e_otax1 sigma p) (l_e_in sigma p o1) (l_e_imagei (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) o1)) (l_e_out sigma p (l_e_in sigma p o1) (l_e_inp sigma p o1)) (l_e_isst1 (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) (l_e_otax1 sigma p) o1) (l_e_isinv (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) (l_e_otax1 sigma p) (l_e_in sigma p o1) (l_e_imagei (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) o1) (l_e_in sigma p o1) (l_e_otax2 sigma p (l_e_in sigma p o1) (l_e_inp sigma p o1)) (l_e_refis sigma (l_e_in sigma p o1)))))). Time Defined. (* constant 271 *) Definition l_e_isinout : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (sp:p s), l_e_is sigma s (l_e_in sigma p (l_e_out sigma p s sp)))))). exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (s:sigma) => (fun (sp:p s) => l_e_ists1 (l_e_ot sigma p) sigma (fun (x:l_e_ot sigma p) => l_e_in sigma p x) (l_e_otax1 sigma p) s (l_e_otax2 sigma p s sp))))). Time Defined. (* constant 272 *) Axiom l_e_pairtype : (forall (sigma:Type), (forall (tau:Type), Type)). (* constant 273 *) Axiom l_e_pair : (forall (sigma:Type), (forall (tau:Type), (forall (s:sigma), (forall (t:tau), l_e_pairtype sigma tau)))). (* constant 274 *) Axiom l_e_first : (forall (sigma:Type), (forall (tau:Type), (forall (p1:l_e_pairtype sigma tau), sigma))). (* constant 275 *) Axiom l_e_second : (forall (sigma:Type), (forall (tau:Type), (forall (p1:l_e_pairtype sigma tau), tau))). (* constant 276 *) Axiom l_e_pairis1 : (forall (sigma:Type), (forall (tau:Type), (forall (p1:l_e_pairtype sigma tau), l_e_is (l_e_pairtype sigma tau) (l_e_pair sigma tau (l_e_first sigma tau p1) (l_e_second sigma tau p1)) p1))). (* constant 277 *) Definition l_e_pairis2 : (forall (sigma:Type), (forall (tau:Type), (forall (p1:l_e_pairtype sigma tau), l_e_is (l_e_pairtype sigma tau) p1 (l_e_pair sigma tau (l_e_first sigma tau p1) (l_e_second sigma tau p1))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (p1:l_e_pairtype sigma tau) => l_e_symis (l_e_pairtype sigma tau) (l_e_pair sigma tau (l_e_first sigma tau p1) (l_e_second sigma tau p1)) p1 (l_e_pairis1 sigma tau p1)))). Time Defined. (* constant 278 *) Axiom l_e_firstis1 : (forall (sigma:Type), (forall (tau:Type), (forall (s:sigma), (forall (t:tau), l_e_is sigma (l_e_first sigma tau (l_e_pair sigma tau s t)) s)))). (* constant 279 *) Definition l_e_firstis2 : (forall (sigma:Type), (forall (tau:Type), (forall (s:sigma), (forall (t:tau), l_e_is sigma s (l_e_first sigma tau (l_e_pair sigma tau s t)))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (s:sigma) => (fun (t:tau) => l_e_symis sigma (l_e_first sigma tau (l_e_pair sigma tau s t)) s (l_e_firstis1 sigma tau s t))))). Time Defined. (* constant 280 *) Axiom l_e_secondis1 : (forall (sigma:Type), (forall (tau:Type), (forall (s:sigma), (forall (t:tau), l_e_is tau (l_e_second sigma tau (l_e_pair sigma tau s t)) t)))). (* constant 281 *) Definition l_e_secondis2 : (forall (sigma:Type), (forall (tau:Type), (forall (s:sigma), (forall (t:tau), l_e_is tau t (l_e_second sigma tau (l_e_pair sigma tau s t)))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (s:sigma) => (fun (t:tau) => l_e_symis tau (l_e_second sigma tau (l_e_pair sigma tau s t)) t (l_e_secondis1 sigma tau s t))))). Time Defined. (* constant 282 *) Definition l_e_ite_prop1 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (z:ksi), Prop))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (z:ksi) => l_and (l_imp a (l_e_is ksi z x)) (l_imp (l_not a) (l_e_is ksi z y))))))). Time Defined. (* constant 283 *) Definition l_e_ite_t1 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_e_ite_prop1 a ksi x y x1), (forall (py1:l_e_ite_prop1 a ksi x y y1), l_imp a (l_e_is ksi x1 x)))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_e_ite_prop1 a ksi x y x1) => (fun (py1:l_e_ite_prop1 a ksi x y y1) => l_ande1 (l_imp a (l_e_is ksi x1 x)) (l_imp (l_not a) (l_e_is ksi x1 y)) px1))))))))). Time Defined. (* constant 284 *) Definition l_e_ite_t2 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_e_ite_prop1 a ksi x y x1), (forall (py1:l_e_ite_prop1 a ksi x y y1), l_e_is ksi x1 x))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_e_ite_prop1 a ksi x y x1) => (fun (py1:l_e_ite_prop1 a ksi x y y1) => l_mp a (l_e_is ksi x1 x) a1 (l_e_ite_t1 a ksi x y a1 x1 y1 px1 py1)))))))))). Time Defined. (* constant 285 *) Definition l_e_ite_t3 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_e_ite_prop1 a ksi x y x1), (forall (py1:l_e_ite_prop1 a ksi x y y1), l_e_is ksi y1 x))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_e_ite_prop1 a ksi x y x1) => (fun (py1:l_e_ite_prop1 a ksi x y y1) => l_e_ite_t2 a ksi x y a1 y1 x1 py1 px1))))))))). Time Defined. (* constant 286 *) Definition l_e_ite_t4 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_e_ite_prop1 a ksi x y x1), (forall (py1:l_e_ite_prop1 a ksi x y y1), l_e_is ksi x1 y1))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_e_ite_prop1 a ksi x y x1) => (fun (py1:l_e_ite_prop1 a ksi x y y1) => l_e_tris2 ksi x1 y1 x (l_e_ite_t2 a ksi x y a1 x1 y1 px1 py1) (l_e_ite_t3 a ksi x y a1 x1 y1 px1 py1)))))))))). Time Defined. (* constant 287 *) Definition l_e_ite_t5 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), l_e_amone ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => (fun (s:ksi) => (fun (t:ksi) => (fun (ps:l_e_ite_prop1 a ksi x y s) => (fun (pt:l_e_ite_prop1 a ksi x y t) => l_e_ite_t4 a ksi x y a1 s t ps pt))))))))). Time Defined. (* constant 288 *) Definition l_e_ite_t6 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), l_imp a (l_e_is ksi x x)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => (fun (x1:a) => l_e_refis ksi x)))))). Time Defined. (* constant 289 *) Definition l_e_ite_t7 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), l_imp (l_not a) (l_e_is ksi x y)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => l_imp_th2 (l_not a) (l_e_is ksi x y) (l_weli a a1)))))). Time Defined. (* constant 290 *) Definition l_e_ite_t8 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), l_e_ite_prop1 a ksi x y x))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => l_andi (l_imp a (l_e_is ksi x x)) (l_imp (l_not a) (l_e_is ksi x y)) (l_e_ite_t6 a ksi x y a1) (l_e_ite_t7 a ksi x y a1)))))). Time Defined. (* constant 291 *) Definition l_e_ite_t9 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), l_some ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => l_somei ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t) x (l_e_ite_t8 a ksi x y a1)))))). Time Defined. (* constant 292 *) Definition l_e_ite_t10 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), l_e_one ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => l_e_onei ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t) (l_e_ite_t5 a ksi x y a1) (l_e_ite_t9 a ksi x y a1)))))). Time Defined. (* constant 293 *) Definition l_e_ite_t11 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_e_ite_prop1 a ksi x y x1), (forall (py1:l_e_ite_prop1 a ksi x y y1), l_imp (l_not a) (l_e_is ksi x1 y)))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_e_ite_prop1 a ksi x y x1) => (fun (py1:l_e_ite_prop1 a ksi x y y1) => l_ande2 (l_imp a (l_e_is ksi x1 x)) (l_imp (l_not a) (l_e_is ksi x1 y)) px1))))))))). Time Defined. (* constant 294 *) Definition l_e_ite_t12 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_e_ite_prop1 a ksi x y x1), (forall (py1:l_e_ite_prop1 a ksi x y y1), l_e_is ksi x1 y))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_e_ite_prop1 a ksi x y x1) => (fun (py1:l_e_ite_prop1 a ksi x y y1) => l_mp (l_not a) (l_e_is ksi x1 y) n (l_e_ite_t11 a ksi x y n x1 y1 px1 py1)))))))))). Time Defined. (* constant 295 *) Definition l_e_ite_t13 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_e_ite_prop1 a ksi x y x1), (forall (py1:l_e_ite_prop1 a ksi x y y1), l_e_is ksi y1 y))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_e_ite_prop1 a ksi x y x1) => (fun (py1:l_e_ite_prop1 a ksi x y y1) => l_e_ite_t12 a ksi x y n y1 x1 py1 px1))))))))). Time Defined. (* constant 296 *) Definition l_e_ite_t14 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_e_ite_prop1 a ksi x y x1), (forall (py1:l_e_ite_prop1 a ksi x y y1), l_e_is ksi x1 y1))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_e_ite_prop1 a ksi x y x1) => (fun (py1:l_e_ite_prop1 a ksi x y y1) => l_e_tris2 ksi x1 y1 y (l_e_ite_t12 a ksi x y n x1 y1 px1 py1) (l_e_ite_t13 a ksi x y n x1 y1 px1 py1)))))))))). Time Defined. (* constant 297 *) Definition l_e_ite_t15 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), l_e_amone ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => (fun (s:ksi) => (fun (t:ksi) => (fun (ps:l_e_ite_prop1 a ksi x y s) => (fun (pt:l_e_ite_prop1 a ksi x y t) => l_e_ite_t14 a ksi x y n s t ps pt))))))))). Time Defined. (* constant 298 *) Definition l_e_ite_t16 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), l_imp (l_not a) (l_e_is ksi y y)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => (fun (x1:l_not a) => l_e_refis ksi y)))))). Time Defined. (* constant 299 *) Definition l_e_ite_t17 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), l_imp a (l_e_is ksi y x)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => l_imp_th2 a (l_e_is ksi y x) n))))). Time Defined. (* constant 300 *) Definition l_e_ite_t18 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), l_e_ite_prop1 a ksi x y y))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => l_andi (l_imp a (l_e_is ksi y x)) (l_imp (l_not a) (l_e_is ksi y y)) (l_e_ite_t17 a ksi x y n) (l_e_ite_t16 a ksi x y n)))))). Time Defined. (* constant 301 *) Definition l_e_ite_t19 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), l_some ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => l_somei ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t) y (l_e_ite_t18 a ksi x y n)))))). Time Defined. (* constant 302 *) Definition l_e_ite_t20 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), l_e_one ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => l_e_onei ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t) (l_e_ite_t15 a ksi x y n) (l_e_ite_t19 a ksi x y n)))))). Time Defined. (* constant 303 *) Definition l_e_ite_t21 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), l_e_one ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => l_imp_th1 a (l_e_one ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t)) (fun (t:a) => l_e_ite_t10 a ksi x y t) (fun (t:l_not a) => l_e_ite_t20 a ksi x y t))))). Time Defined. (* constant 304 *) Definition l_e_ite : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), ksi)))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => l_e_ind ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t) (l_e_ite_t21 a ksi x y))))). Time Defined. (* constant 305 *) Definition l_e_ite_t22 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), l_e_ite_prop1 a ksi x y (l_e_ite a ksi x y))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => l_e_oneax ksi (fun (t:ksi) => l_e_ite_prop1 a ksi x y t) (l_e_ite_t21 a ksi x y))))). Time Defined. (* constant 306 *) Definition l_e_ite_t23 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), l_imp a (l_e_is ksi (l_e_ite a ksi x y) x))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => l_ande1 (l_imp a (l_e_is ksi (l_e_ite a ksi x y) x)) (l_imp (l_not a) (l_e_is ksi (l_e_ite a ksi x y) y)) (l_e_ite_t22 a ksi x y))))). Time Defined. (* constant 307 *) Definition l_e_ite_t24 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), l_imp (l_not a) (l_e_is ksi (l_e_ite a ksi x y) y))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => l_ande2 (l_imp a (l_e_is ksi (l_e_ite a ksi x y) x)) (l_imp (l_not a) (l_e_is ksi (l_e_ite a ksi x y) y)) (l_e_ite_t22 a ksi x y))))). Time Defined. (* constant 308 *) Definition l_e_itet : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (a1:a), l_e_is ksi (l_e_ite a ksi x y) x))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => l_mp a (l_e_is ksi (l_e_ite a ksi x y) x) a1 (l_e_ite_t23 a ksi x y)))))). Time Defined. (* constant 309 *) Definition l_e_itef : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (n:l_not a), l_e_is ksi (l_e_ite a ksi x y) y))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (n:l_not a) => l_mp (l_not a) (l_e_is ksi (l_e_ite a ksi x y) y) n (l_e_ite_t24 a ksi x y)))))). Time Defined. (* constant 310 *) Definition l_e_wissel_wa : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), sigma)))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => l_e_ite (l_e_is sigma s s0) sigma t0 s)))). Time Defined. (* constant 311 *) Definition l_e_wissel_t1 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s s0), l_e_is sigma (l_e_wissel_wa sigma s0 t0 s) t0))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s s0) => l_e_itet (l_e_is sigma s s0) sigma t0 s i))))). Time Defined. (* constant 312 *) Definition l_e_wissel_t2 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (n:l_not (l_e_is sigma s s0)), l_e_is sigma (l_e_wissel_wa sigma s0 t0 s) s))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_is sigma s s0)) => l_e_itef (l_e_is sigma s s0) sigma t0 s n))))). Time Defined. (* constant 313 *) Definition l_e_wissel_wb : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), sigma)))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => l_e_ite (l_e_is sigma s t0) sigma s0 (l_e_wissel_wa sigma s0 t0 s))))). Time Defined. (* constant 314 *) Definition l_e_wissel_t3 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s t0), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) s0))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s t0) => l_e_itet (l_e_is sigma s t0) sigma s0 (l_e_wissel_wa sigma s0 t0 s) i))))). Time Defined. (* constant 315 *) Definition l_e_wissel_t4 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (n:l_not (l_e_is sigma s t0)), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wa sigma s0 t0 s)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_is sigma s t0)) => l_e_itef (l_e_is sigma s t0) sigma s0 (l_e_wissel_wa sigma s0 t0 s) n))))). Time Defined. (* constant 316 *) Definition l_e_wissel_t5 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s s0), (forall (j:l_e_is sigma s0 t0), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) t0)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s s0) => (fun (j:l_e_is sigma s0 t0) => l_e_tris sigma (l_e_wissel_wb sigma s0 t0 s) s0 t0 (l_e_wissel_t3 sigma s0 t0 s (l_e_tris sigma s s0 t0 i j)) j)))))). Time Defined. (* constant 317 *) Definition l_e_wissel_t6 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s s0), (forall (n:l_not (l_e_is sigma s0 t0)), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) t0)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s s0) => (fun (n:l_not (l_e_is sigma s0 t0)) => l_e_tris sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wa sigma s0 t0 s) t0 (l_e_wissel_t4 sigma s0 t0 s (l_e_notis_th2 sigma s0 t0 s n i)) (l_e_wissel_t1 sigma s0 t0 s i))))))). Time Defined. (* constant 318 *) Definition l_e_wissel_t7 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s s0), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) t0))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s s0) => l_imp_th1 (l_e_is sigma s0 t0) (l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) t0) (fun (t:l_e_is sigma s0 t0) => l_e_wissel_t5 sigma s0 t0 s i t) (fun (t:l_not (l_e_is sigma s0 t0)) => l_e_wissel_t6 sigma s0 t0 s i t)))))). Time Defined. (* constant 319 *) Definition l_e_wissel_t8 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (n:l_not (l_e_is sigma s s0)), (forall (o:l_not (l_e_is sigma s t0)), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) s)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_is sigma s s0)) => (fun (o:l_not (l_e_is sigma s t0)) => l_e_tris sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wa sigma s0 t0 s) s (l_e_wissel_t4 sigma s0 t0 s o) (l_e_wissel_t2 sigma s0 t0 s n))))))). Time Defined. (* constant 320 *) Definition l_e_wissel : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (x:sigma), sigma)))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (x:sigma) => l_e_wissel_wb sigma s0 t0 x)))). Time Defined. (* constant 321 *) Definition l_e_iswissel1 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s s0), l_e_is sigma (l_e_wissel sigma s0 t0 s) t0))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s s0) => l_e_wissel_t7 sigma s0 t0 s i))))). Time Defined. (* constant 322 *) Definition l_e_iswissel2 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s t0), l_e_is sigma (l_e_wissel sigma s0 t0 s) s0))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s t0) => l_e_wissel_t3 sigma s0 t0 s i))))). Time Defined. (* constant 323 *) Definition l_e_iswissel3 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (n:l_not (l_e_is sigma s s0)), (forall (o:l_not (l_e_is sigma s t0)), l_e_is sigma (l_e_wissel sigma s0 t0 s) s)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_is sigma s s0)) => (fun (o:l_not (l_e_is sigma s t0)) => l_e_wissel_t8 sigma s0 t0 s n o)))))). Time Defined. (* constant 324 *) Definition l_e_wissel_t9 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), l_not (l_e_is sigma t s0))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => l_e_symnotis sigma s0 t (l_e_notis_th1 sigma s t s0 n j))))))))). Time Defined. (* constant 325 *) Definition l_e_wissel_t10 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), (forall (k:l_e_is sigma s0 t0), l_not (l_e_is sigma t t0)))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => (fun (k:l_e_is sigma s0 t0) => l_e_notis_th3 sigma t s0 t0 (l_e_wissel_t9 sigma s0 t0 s t i n j) k))))))))). Time Defined. (* constant 326 *) Definition l_e_wissel_t11 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), (forall (k:l_e_is sigma s0 t0), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) t))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => (fun (k:l_e_is sigma s0 t0) => l_e_tris sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t) t i (l_e_wissel_t8 sigma s0 t0 t (l_e_wissel_t9 sigma s0 t0 s t i n j) (l_e_wissel_t10 sigma s0 t0 s t i n j k))))))))))). Time Defined. (* constant 327 *) Definition l_e_wissel_t12 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), (forall (k:l_e_is sigma s0 t0), l_con))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => (fun (k:l_e_is sigma s0 t0) => l_e_wissel_t10 sigma s0 t0 s t i n j k (l_e_tris1 sigma t t0 (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_t11 sigma s0 t0 s t i n j k) (l_e_wissel_t7 sigma s0 t0 s j))))))))))). Time Defined. (* constant 328 *) Definition l_e_wissel_t13 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), l_not (l_e_is sigma s0 t0))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => (fun (v:l_e_is sigma s0 t0) => l_e_wissel_t12 sigma s0 t0 s t i n j v))))))))). Time Defined. (* constant 329 *) Definition l_e_wissel_t14 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), (forall (k:l_e_is sigma t t0), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) s0))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => (fun (k:l_e_is sigma t t0) => l_e_tris sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t) s0 i (l_e_wissel_t3 sigma s0 t0 t k)))))))))). Time Defined. (* constant 330 *) Definition l_e_wissel_t15 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), (forall (k:l_e_is sigma t t0), l_con))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => (fun (k:l_e_is sigma t t0) => l_e_wissel_t12 sigma s0 t0 s t i n j (l_e_tris1 sigma s0 t0 (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_t14 sigma s0 t0 s t i n j k) (l_e_wissel_t7 sigma s0 t0 s j))))))))))). Time Defined. (* constant 331 *) Definition l_e_wissel_t16 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), l_not (l_e_is sigma t t0))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => (fun (v:l_e_is sigma t t0) => l_e_wissel_t15 sigma s0 t0 s t i n j v))))))))). Time Defined. (* constant 332 *) Definition l_e_wissel_t17 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) t)))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => l_e_tris sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t) t i (l_e_wissel_t8 sigma s0 t0 t (l_e_wissel_t9 sigma s0 t0 s t i n j) (l_e_wissel_t16 sigma s0 t0 s t i n j)))))))))). Time Defined. (* constant 333 *) Definition l_e_wissel_t18 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s s0), l_con)))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s s0) => l_e_wissel_t15 sigma s0 t0 s t i n j (l_e_tris1 sigma t t0 (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_t17 sigma s0 t0 s t i n j) (l_e_wissel_t7 sigma s0 t0 s j)))))))))). Time Defined. (* constant 334 *) Definition l_e_wissel_t19 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), l_not (l_e_is sigma s s0)))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (v:l_e_is sigma s s0) => l_e_wissel_t18 sigma s0 t0 s t i n v)))))))). Time Defined. (* constant 335 *) Definition l_e_wissel_t20 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), l_not (l_e_is sigma t s0)))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => l_e_wissel_t19 sigma s0 t0 t s (l_e_symis sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t) i) (l_e_symnotis sigma s t n)))))))). Time Defined. (* constant 336 *) Definition l_e_wissel_t21 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s t0), l_not (l_e_is sigma t t0))))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s t0) => l_e_symnotis sigma t0 t (l_e_notis_th1 sigma s t t0 n j))))))))). Time Defined. (* constant 337 *) Definition l_e_wissel_t22 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s t0), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) t)))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s t0) => l_e_tris sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t) t i (l_e_wissel_t8 sigma s0 t0 t (l_e_wissel_t20 sigma s0 t0 s t i n) (l_e_wissel_t21 sigma s0 t0 s t i n j)))))))))). Time Defined. (* constant 338 *) Definition l_e_wissel_t23 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), (forall (j:l_e_is sigma s t0), l_con)))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (j:l_e_is sigma s t0) => l_e_wissel_t20 sigma s0 t0 s t i n (l_e_tris1 sigma t s0 (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_t22 sigma s0 t0 s t i n j) (l_e_wissel_t3 sigma s0 t0 s j)))))))))). Time Defined. (* constant 339 *) Definition l_e_wissel_t24 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), l_not (l_e_is sigma s t0)))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => (fun (v:l_e_is sigma s t0) => l_e_wissel_t23 sigma s0 t0 s t i n v)))))))). Time Defined. (* constant 340 *) Definition l_e_wissel_t25 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), l_not (l_e_is sigma t t0)))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => l_e_wissel_t24 sigma s0 t0 t s (l_e_symis sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t) i) (l_e_symnotis sigma s t n)))))))). Time Defined. (* constant 341 *) Definition l_e_wissel_t26 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) t))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => l_e_tris sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t) t i (l_e_wissel_t8 sigma s0 t0 t (l_e_wissel_t20 sigma s0 t0 s t i n) (l_e_wissel_t25 sigma s0 t0 s t i n))))))))). Time Defined. (* constant 342 *) Definition l_e_wissel_t27 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), (forall (n:l_not (l_e_is sigma s t)), l_con))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => (fun (n:l_not (l_e_is sigma s t)) => n (l_e_tris1 sigma s t (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_t8 sigma s0 t0 s (l_e_wissel_t19 sigma s0 t0 s t i n) (l_e_wissel_t24 sigma s0 t0 s t i n)) (l_e_wissel_t26 sigma s0 t0 s t i n))))))))). Time Defined. (* constant 343 *) Definition l_e_wissel_t28 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (t:sigma), (forall (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)), l_e_is sigma s t)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (t:sigma) => (fun (i:l_e_is sigma (l_e_wissel_wb sigma s0 t0 s) (l_e_wissel_wb sigma s0 t0 t)) => l_et (l_e_is sigma s t) (fun (v:l_not (l_e_is sigma s t)) => l_e_wissel_t27 sigma s0 t0 s t i v))))))). Time Defined. (* constant 344 *) Definition l_e_wissel_th1 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), l_e_injective sigma sigma (l_e_wissel sigma s0 t0)))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (x:sigma) => (fun (y:sigma) => (fun (v:l_e_is sigma (l_e_wissel_wb sigma s0 t0 x) (l_e_wissel_wb sigma s0 t0 y)) => l_e_wissel_t28 sigma s0 t0 x y v)))))). Time Defined. (* constant 345 *) Definition l_e_wissel_t29 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s s0), l_e_is sigma s (l_e_wissel_wb sigma s0 t0 t0)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s s0) => l_e_tris2 sigma s (l_e_wissel_wb sigma s0 t0 t0) s0 i (l_e_wissel_t3 sigma s0 t0 t0 (l_e_refis sigma t0))))))). Time Defined. (* constant 346 *) Definition l_e_wissel_t30 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s s0), l_e_image sigma sigma (l_e_wissel sigma s0 t0) s))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s s0) => l_somei sigma (fun (x:sigma) => l_e_is sigma s (l_e_wissel_wb sigma s0 t0 x)) t0 (l_e_wissel_t29 sigma s0 t0 s i)))))). Time Defined. (* constant 347 *) Definition l_e_wissel_t31 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s t0), l_e_is sigma s (l_e_wissel_wb sigma s0 t0 s0)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s t0) => l_e_tris2 sigma s (l_e_wissel_wb sigma s0 t0 s0) t0 i (l_e_wissel_t7 sigma s0 t0 s0 (l_e_refis sigma s0))))))). Time Defined. (* constant 348 *) Definition l_e_wissel_t32 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s t0), l_e_image sigma sigma (l_e_wissel sigma s0 t0) s))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s t0) => l_somei sigma (fun (x:sigma) => l_e_is sigma s (l_e_wissel_wb sigma s0 t0 x)) s0 (l_e_wissel_t31 sigma s0 t0 s i)))))). Time Defined. (* constant 349 *) Definition l_e_wissel_t33 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (n:l_not (l_e_is sigma s s0)), (forall (o:l_not (l_e_is sigma s t0)), l_e_is sigma s (l_e_wissel_wb sigma s0 t0 s))))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_is sigma s s0)) => (fun (o:l_not (l_e_is sigma s t0)) => l_e_symis sigma (l_e_wissel_wb sigma s0 t0 s) s (l_e_wissel_t8 sigma s0 t0 s n o))))))). Time Defined. (* constant 350 *) Definition l_e_wissel_t34 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (n:l_not (l_e_is sigma s s0)), (forall (o:l_not (l_e_is sigma s t0)), l_e_image sigma sigma (l_e_wissel sigma s0 t0) s)))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_is sigma s s0)) => (fun (o:l_not (l_e_is sigma s t0)) => l_somei sigma (fun (x:sigma) => l_e_is sigma s (l_e_wissel_wb sigma s0 t0 x)) s (l_e_wissel_t33 sigma s0 t0 s n o))))))). Time Defined. (* constant 351 *) Definition l_e_wissel_t35 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (n:l_not (l_e_is sigma s s0)), l_e_image sigma sigma (l_e_wissel sigma s0 t0) s))))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_is sigma s s0)) => l_imp_th1 (l_e_is sigma s t0) (l_e_image sigma sigma (l_e_wissel sigma s0 t0) s) (fun (v:l_e_is sigma s t0) => l_e_wissel_t32 sigma s0 t0 s v) (fun (v:l_not (l_e_is sigma s t0)) => l_e_wissel_t34 sigma s0 t0 s n v)))))). Time Defined. (* constant 352 *) Definition l_e_wissel_t36 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), l_e_image sigma sigma (l_e_wissel sigma s0 t0) s)))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => l_imp_th1 (l_e_is sigma s s0) (l_e_image sigma sigma (l_e_wissel sigma s0 t0) s) (fun (v:l_e_is sigma s s0) => l_e_wissel_t30 sigma s0 t0 s v) (fun (v:l_not (l_e_is sigma s s0)) => l_e_wissel_t35 sigma s0 t0 s v))))). Time Defined. (* constant 353 *) Definition l_e_wissel_th2 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), l_e_surjective sigma sigma (l_e_wissel sigma s0 t0)))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (x:sigma) => l_e_wissel_t36 sigma s0 t0 x)))). Time Defined. (* constant 354 *) Definition l_e_wissel_th3 : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), l_e_bijective sigma sigma (l_e_wissel sigma s0 t0)))). exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => l_andi (l_e_injective sigma sigma (l_e_wissel sigma s0 t0)) (l_e_surjective sigma sigma (l_e_wissel sigma s0 t0)) (l_e_wissel_th1 sigma s0 t0) (l_e_wissel_th2 sigma s0 t0)))). Time Defined. (* constant 355 *) Definition l_e_changef : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s0:sigma), (forall (t0:sigma), (forall (x:sigma), tau)))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (x:sigma) => f (l_e_wissel sigma s0 t0 x))))))). Time Defined. (* constant 356 *) Definition l_e_changef1 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s s0), l_e_is tau (l_e_changef sigma tau f s0 t0 s) (f t0)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s s0) => l_e_isf sigma tau f (l_e_wissel sigma s0 t0 s) t0 (l_e_iswissel1 sigma s0 t0 s i)))))))). Time Defined. (* constant 357 *) Definition l_e_changef2 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (i:l_e_is sigma s t0), l_e_is tau (l_e_changef sigma tau f s0 t0 s) (f s0)))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (i:l_e_is sigma s t0) => l_e_isf sigma tau f (l_e_wissel sigma s0 t0 s) s0 (l_e_iswissel2 sigma s0 t0 s i)))))))). Time Defined. (* constant 358 *) Definition l_e_changef3 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), (forall (n:l_not (l_e_is sigma s s0)), (forall (o:l_not (l_e_is sigma s t0)), l_e_is tau (l_e_changef sigma tau f s0 t0 s) (f s))))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_is sigma s s0)) => (fun (o:l_not (l_e_is sigma s t0)) => l_e_isf sigma tau f (l_e_wissel sigma s0 t0 s) s (l_e_iswissel3 sigma s0 t0 s n o))))))))). Time Defined. (* constant 359 *) Definition l_e_wissel_th4 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s0:sigma), (forall (t0:sigma), (forall (i:l_e_injective sigma tau f), l_e_injective sigma tau (l_e_changef sigma tau f s0 t0))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (i:l_e_injective sigma tau f) => l_e_inj_th4 sigma sigma tau (l_e_wissel sigma s0 t0) f (l_e_wissel_th1 sigma s0 t0) i)))))). Time Defined. (* constant 360 *) Definition l_e_wissel_th5 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s0:sigma), (forall (t0:sigma), (forall (s:l_e_surjective sigma tau f), l_e_surjective sigma tau (l_e_changef sigma tau f s0 t0))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (s:l_e_surjective sigma tau f) => l_e_surj_th1 sigma sigma tau (l_e_wissel sigma s0 t0) f (l_e_wissel_th2 sigma s0 t0) s)))))). Time Defined. (* constant 361 *) Definition l_e_wissel_th6 : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s0:sigma), (forall (t0:sigma), (forall (b:l_e_bijective sigma tau f), l_e_bijective sigma tau (l_e_changef sigma tau f s0 t0))))))). exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (b:l_e_bijective sigma tau f) => l_e_bij_th1 sigma sigma tau (l_e_wissel sigma s0 t0) f (l_e_wissel_th3 sigma s0 t0) b)))))). Time Defined. (* constant 362 *) Definition l_r_imp : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), Prop)). exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (forall (x:a), b x))). Time Defined. (* constant 363 *) Definition l_r_mp : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), (forall (a1:a), (forall (i:l_r_imp a b), b a1)))). exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (fun (a1:a) => (fun (i:l_r_imp a b) => i a1)))). Time Defined. (* constant 364 *) Definition l_r_imp_th2 : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), (forall (n:l_not a), l_r_imp a b))). exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (fun (n:l_not a) => (fun (x:a) => l_cone (b x) (l_mp a l_con x n))))). Time Defined. (* constant 365 *) Definition l_r_ec : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), Prop)). exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (forall (x:a), l_not (b x)))). Time Defined. (* constant 366 *) Definition l_r_eci1 : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), (forall (n:l_not a), l_r_ec a b))). exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (fun (n:l_not a) => (fun (x:a) => l_cone (l_not (b x)) (l_mp a l_con x n))))). Time Defined. (* constant 367 *) Definition l_r_ande2 : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), (forall (a1:l_and a (l_r_imp a b)), b (l_ande1 a (l_r_imp a b) a1)))). exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (fun (a1:l_and a (l_r_imp a b)) => l_ande2 a (l_r_imp a b) a1 (l_ande1 a (l_r_imp a b) a1)))). Time Defined. (* constant 368 *) Definition l_r_ite_is : (forall (a:Prop), (forall (ksi:Type), (forall (x1:ksi), (forall (y1:ksi), Prop)))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x1:ksi) => (fun (y1:ksi) => l_e_is ksi x1 y1)))). Time Defined. (* constant 369 *) Definition l_r_ite_prop1 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (z:ksi), Prop))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (z:ksi) => l_and (l_r_imp a (fun (t:a) => l_r_ite_is a ksi z (x t))) (l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi z (y t)))))))))). Time Defined. (* constant 370 *) Definition l_r_ite_t1 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_r_ite_prop1 a ksi x y i j x1), (forall (py1:l_r_ite_prop1 a ksi x y i j y1), l_r_imp a (fun (t:a) => l_r_ite_is a ksi x1 (x t))))))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_r_ite_prop1 a ksi x y i j x1) => (fun (py1:l_r_ite_prop1 a ksi x y i j y1) => l_ande1 (l_r_imp a (fun (t:a) => l_r_ite_is a ksi x1 (x t))) (l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi x1 (y t))) px1))))))))))). Time Defined. (* constant 371 *) Definition l_r_ite_t2 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_r_ite_prop1 a ksi x y i j x1), (forall (py1:l_r_ite_prop1 a ksi x y i j y1), l_r_ite_is a ksi x1 (x a1)))))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_r_ite_prop1 a ksi x y i j x1) => (fun (py1:l_r_ite_prop1 a ksi x y i j y1) => l_r_mp a (fun (t:a) => l_r_ite_is a ksi x1 (x t)) a1 (l_r_ite_t1 a ksi x y i j a1 x1 y1 px1 py1)))))))))))). Time Defined. (* constant 372 *) Definition l_r_ite_t3 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_r_ite_prop1 a ksi x y i j x1), (forall (py1:l_r_ite_prop1 a ksi x y i j y1), l_r_ite_is a ksi y1 (x a1)))))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_r_ite_prop1 a ksi x y i j x1) => (fun (py1:l_r_ite_prop1 a ksi x y i j y1) => l_r_ite_t2 a ksi x y i j a1 y1 x1 py1 px1))))))))))). Time Defined. (* constant 373 *) Definition l_r_ite_t4 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_r_ite_prop1 a ksi x y i j x1), (forall (py1:l_r_ite_prop1 a ksi x y i j y1), l_r_ite_is a ksi x1 y1))))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_r_ite_prop1 a ksi x y i j x1) => (fun (py1:l_r_ite_prop1 a ksi x y i j y1) => l_e_tris2 ksi x1 y1 (x a1) (l_r_ite_t2 a ksi x y i j a1 x1 y1 px1 py1) (l_r_ite_t3 a ksi x y i j a1 x1 y1 px1 py1)))))))))))). Time Defined. (* constant 374 *) Definition l_r_ite_t5 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), l_e_amone ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => (fun (s:ksi) => (fun (t:ksi) => (fun (ps:l_r_ite_prop1 a ksi x y i j s) => (fun (pt:l_r_ite_prop1 a ksi x y i j t) => l_r_ite_t4 a ksi x y i j a1 s t ps pt))))))))))). Time Defined. (* constant 375 *) Definition l_r_ite_t6 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), l_r_imp a (fun (t:a) => l_r_ite_is a ksi (x a1) (x t))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => i a1))))))). Time Defined. (* constant 376 *) Definition l_r_ite_t7 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (x a1) (y t))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => l_r_imp_th2 (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (x a1) (y t)) (l_weli a a1)))))))). Time Defined. (* constant 377 *) Definition l_r_ite_t8 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), l_r_ite_prop1 a ksi x y i j (x a1)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => l_andi (l_r_imp a (fun (t:a) => l_r_ite_is a ksi (x a1) (x t))) (l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (x a1) (y t))) (l_r_ite_t6 a ksi x y i j a1) (l_r_ite_t7 a ksi x y i j a1)))))))). Time Defined. (* constant 378 *) Definition l_r_ite_t9 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), l_some ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => l_somei ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t) (x a1) (l_r_ite_t8 a ksi x y i j a1)))))))). Time Defined. (* constant 379 *) Definition l_r_ite_t10 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), l_e_one ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => l_e_onei ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t) (l_r_ite_t5 a ksi x y i j a1) (l_r_ite_t9 a ksi x y i j a1)))))))). Time Defined. (* constant 380 *) Definition l_r_ite_t11 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_r_ite_prop1 a ksi x y i j x1), (forall (py1:l_r_ite_prop1 a ksi x y i j y1), l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi x1 (y t))))))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_r_ite_prop1 a ksi x y i j x1) => (fun (py1:l_r_ite_prop1 a ksi x y i j y1) => l_ande2 (l_r_imp a (fun (t:a) => l_r_ite_is a ksi x1 (x t))) (l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi x1 (y t))) px1))))))))))). Time Defined. (* constant 381 *) Definition l_r_ite_t12 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_r_ite_prop1 a ksi x y i j x1), (forall (py1:l_r_ite_prop1 a ksi x y i j y1), l_r_ite_is a ksi x1 (y n)))))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_r_ite_prop1 a ksi x y i j x1) => (fun (py1:l_r_ite_prop1 a ksi x y i j y1) => l_r_mp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi x1 (y t)) n (l_r_ite_t11 a ksi x y i j n x1 y1 px1 py1)))))))))))). Time Defined. (* constant 382 *) Definition l_r_ite_t13 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_r_ite_prop1 a ksi x y i j x1), (forall (py1:l_r_ite_prop1 a ksi x y i j y1), l_r_ite_is a ksi y1 (y n)))))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_r_ite_prop1 a ksi x y i j x1) => (fun (py1:l_r_ite_prop1 a ksi x y i j y1) => l_r_ite_t12 a ksi x y i j n y1 x1 py1 px1))))))))))). Time Defined. (* constant 383 *) Definition l_r_ite_t14 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), (forall (x1:ksi), (forall (y1:ksi), (forall (px1:l_r_ite_prop1 a ksi x y i j x1), (forall (py1:l_r_ite_prop1 a ksi x y i j y1), l_r_ite_is a ksi x1 y1))))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => (fun (x1:ksi) => (fun (y1:ksi) => (fun (px1:l_r_ite_prop1 a ksi x y i j x1) => (fun (py1:l_r_ite_prop1 a ksi x y i j y1) => l_e_tris2 ksi x1 y1 (y n) (l_r_ite_t12 a ksi x y i j n x1 y1 px1 py1) (l_r_ite_t13 a ksi x y i j n x1 y1 px1 py1)))))))))))). Time Defined. (* constant 384 *) Definition l_r_ite_t15 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), l_e_amone ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => (fun (s:ksi) => (fun (t:ksi) => (fun (ps:l_r_ite_prop1 a ksi x y i j s) => (fun (pt:l_r_ite_prop1 a ksi x y i j t) => l_r_ite_t14 a ksi x y i j n s t ps pt))))))))))). Time Defined. (* constant 385 *) Definition l_r_ite_t16 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (y n) (y t))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => j n))))))). Time Defined. (* constant 386 *) Definition l_r_ite_t17 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), l_r_imp a (fun (t:a) => l_r_ite_is a ksi (y n) (x t))))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => l_r_imp_th2 a (fun (t:a) => l_r_ite_is a ksi (y n) (x t)) n))))))). Time Defined. (* constant 387 *) Definition l_r_ite_t18 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), l_r_ite_prop1 a ksi x y i j (y n)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => l_andi (l_r_imp a (fun (t:a) => l_r_ite_is a ksi (y n) (x t))) (l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (y n) (y t))) (l_r_ite_t17 a ksi x y i j n) (l_r_ite_t16 a ksi x y i j n)))))))). Time Defined. (* constant 388 *) Definition l_r_ite_t19 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), l_some ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => l_somei ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t) (y n) (l_r_ite_t18 a ksi x y i j n)))))))). Time Defined. (* constant 389 *) Definition l_r_ite_t20 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), l_e_one ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => l_e_onei ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t) (l_r_ite_t15 a ksi x y i j n) (l_r_ite_t19 a ksi x y i j n)))))))). Time Defined. (* constant 390 *) Definition l_r_ite_t21 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), l_e_one ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => l_imp_th1 a (l_e_one ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t)) (fun (t:a) => l_r_ite_t10 a ksi x y i j t) (fun (t:l_not a) => l_r_ite_t20 a ksi x y i j t))))))). Time Defined. (* constant 391 *) Definition l_r_ite : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), ksi)))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => l_e_ind ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t) (l_r_ite_t21 a ksi x y i j))))))). Time Defined. (* constant 392 *) Definition l_r_ite_t22 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), l_r_ite_prop1 a ksi x y i j (l_r_ite a ksi x y i j))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => l_e_oneax ksi (fun (t:ksi) => l_r_ite_prop1 a ksi x y i j t) (l_r_ite_t21 a ksi x y i j))))))). Time Defined. (* constant 393 *) Definition l_r_ite_t23 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), l_r_imp a (fun (t:a) => l_r_ite_is a ksi (l_r_ite a ksi x y i j) (x t)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => l_ande1 (l_r_imp a (fun (t:a) => l_r_ite_is a ksi (l_r_ite a ksi x y i j) (x t))) (l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (l_r_ite a ksi x y i j) (y t))) (l_r_ite_t22 a ksi x y i j))))))). Time Defined. (* constant 394 *) Definition l_r_ite_t24 : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (l_r_ite a ksi x y i j) (y t)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => l_ande2 (l_r_imp a (fun (t:a) => l_r_ite_is a ksi (l_r_ite a ksi x y i j) (x t))) (l_r_imp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (l_r_ite a ksi x y i j) (y t))) (l_r_ite_t22 a ksi x y i j))))))). Time Defined. (* constant 395 *) Definition l_r_itet : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (a1:a), l_r_ite_is a ksi (l_r_ite a ksi x y i j) (x a1)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (a1:a) => l_r_mp a (fun (t:a) => l_r_ite_is a ksi (l_r_ite a ksi x y i j) (x t)) a1 (l_r_ite_t23 a ksi x y i j)))))))). Time Defined. (* constant 396 *) Definition l_r_itef : (forall (a:Prop), (forall (ksi:Type), (forall (x:(forall (t:a), ksi)), (forall (y:(forall (t:l_not a), ksi)), (forall (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))), (forall (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))), (forall (n:l_not a), l_r_ite_is a ksi (l_r_ite a ksi x y i j) (y n)))))))). exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:(forall (t:a), ksi)) => (fun (y:(forall (t:l_not a), ksi)) => (fun (i:(forall (t:a), (forall (u:a), l_r_ite_is a ksi (x t) (x u)))) => (fun (j:(forall (t:l_not a), (forall (u:l_not a), l_r_ite_is a ksi (y t) (y u)))) => (fun (n:l_not a) => l_r_mp (l_not a) (fun (t:l_not a) => l_r_ite_is a ksi (l_r_ite a ksi x y i j) (y t)) n (l_r_ite_t24 a ksi x y i j)))))))). Time Defined. (* constant 397 *) Axiom l_e_st_set : (forall (sigma:Type), Type). (* constant 398 *) Axiom l_e_st_esti : (forall (sigma:Type), (forall (s:sigma), (forall (s0:l_e_st_set sigma), Prop))). (* constant 399 *) Axiom l_e_st_setof : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), l_e_st_set sigma)). (* constant 400 *) Axiom l_e_st_estii : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (sp:p s), l_e_st_esti sigma s (l_e_st_setof sigma p))))). (* constant 401 *) Axiom l_e_st_estie : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (e:l_e_st_esti sigma s (l_e_st_setof sigma p)), p s)))). (* constant 402 *) Definition l_e_st_empty : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), Prop)). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => l_none sigma (fun (x:sigma) => l_e_st_esti sigma x s0))). Time Defined. (* constant 403 *) Definition l_e_st_nonempty : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), Prop)). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => l_some sigma (fun (x:sigma) => l_e_st_esti sigma x s0))). Time Defined. (* constant 404 *) Definition l_e_st_emptyi : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (n:(forall (x:sigma), l_not (l_e_st_esti sigma x s0))), l_e_st_empty sigma s0))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (n:(forall (x:sigma), l_not (l_e_st_esti sigma x s0))) => n))). Time Defined. (* constant 405 *) Definition l_e_st_emptye : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (e:l_e_st_empty sigma s0), (forall (s:sigma), l_not (l_e_st_esti sigma s s0))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (e:l_e_st_empty sigma s0) => (fun (s:sigma) => e s)))). Time Defined. (* constant 406 *) Definition l_e_st_nonemptyi : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), l_e_st_nonempty sigma s0)))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => l_somei sigma (fun (x:sigma) => l_e_st_esti sigma x s0) s ses0)))). Time Defined. (* constant 407 *) Definition l_e_st_nonemptyapp : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (n:l_e_st_nonempty sigma s0), (forall (x:Prop), (forall (x1:(forall (y:sigma), (forall (z:l_e_st_esti sigma y s0), x))), x))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (n:l_e_st_nonempty sigma s0) => (fun (x:Prop) => (fun (x1:(forall (y:sigma), (forall (z:l_e_st_esti sigma y s0), x))) => l_someapp sigma (fun (y:sigma) => l_e_st_esti sigma y s0) n x x1))))). Time Defined. (* constant 408 *) Definition l_e_st_incl : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), Prop))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => l_all sigma (fun (x:sigma) => l_imp (l_e_st_esti sigma x s0) (l_e_st_esti sigma x t0))))). Time Defined. (* constant 409 *) Definition l_e_st_incli : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (e:(forall (x:sigma), (forall (y:l_e_st_esti sigma x s0), l_e_st_esti sigma x t0))), l_e_st_incl sigma s0 t0)))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (e:(forall (x:sigma), (forall (y:l_e_st_esti sigma x s0), l_e_st_esti sigma x t0))) => e)))). Time Defined. (* constant 410 *) Definition l_e_st_incle : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (i:l_e_st_incl sigma s0 t0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), l_e_st_esti sigma s t0)))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (i:l_e_st_incl sigma s0 t0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => i s ses0)))))). Time Defined. (* constant 411 *) Definition l_e_st_refincl : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), l_e_st_incl sigma s0 s0)). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (x:sigma) => (fun (y:l_e_st_esti sigma x s0) => y)))). Time Defined. (* constant 412 *) Definition l_e_st_disj : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), Prop))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => l_all sigma (fun (x:sigma) => l_ec (l_e_st_esti sigma x s0) (l_e_st_esti sigma x t0))))). Time Defined. (* constant 413 *) Definition l_e_st_disji1 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (n:(forall (x:sigma), (forall (y:l_e_st_esti sigma x s0), l_not (l_e_st_esti sigma x t0)))), l_e_st_disj sigma s0 t0)))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (n:(forall (x:sigma), (forall (y:l_e_st_esti sigma x s0), l_not (l_e_st_esti sigma x t0)))) => n)))). Time Defined. (* constant 414 *) Definition l_e_st_disji2 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (n:(forall (x:sigma), (forall (y:l_e_st_esti sigma x t0), l_not (l_e_st_esti sigma x s0)))), l_e_st_disj sigma s0 t0)))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (n:(forall (x:sigma), (forall (y:l_e_st_esti sigma x t0), l_not (l_e_st_esti sigma x s0)))) => (fun (x:sigma) => l_ec_th2 (l_e_st_esti sigma x s0) (l_e_st_esti sigma x t0) (n x)))))). Time Defined. (* constant 415 *) Definition l_e_st_disje1 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (d:l_e_st_disj sigma s0 t0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), l_not (l_e_st_esti sigma s t0))))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (d:l_e_st_disj sigma s0 t0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => l_ece1 (l_e_st_esti sigma s s0) (l_e_st_esti sigma s t0) (d s) ses0)))))). Time Defined. (* constant 416 *) Definition l_e_st_disje2 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (d:l_e_st_disj sigma s0 t0), (forall (s:sigma), (forall (set0:l_e_st_esti sigma s t0), l_not (l_e_st_esti sigma s s0))))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (d:l_e_st_disj sigma s0 t0) => (fun (s:sigma) => (fun (set0:l_e_st_esti sigma s t0) => l_ece2 (l_e_st_esti sigma s s0) (l_e_st_esti sigma s t0) (d s) set0)))))). Time Defined. (* constant 417 *) Definition l_e_st_symdisj : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (d:l_e_st_disj sigma s0 t0), l_e_st_disj sigma t0 s0)))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (d:l_e_st_disj sigma s0 t0) => (fun (x:sigma) => (fun (y:l_e_st_esti sigma x t0) => l_e_st_disje2 sigma s0 t0 d x y)))))). Time Defined. (* constant 418 *) Definition l_e_st_disj_th1 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (set0:l_e_st_esti sigma s t0), l_not (l_e_st_disj sigma s0 t0))))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (set0:l_e_st_esti sigma s t0) => l_all_th1 sigma (fun (x:sigma) => l_ec (l_e_st_esti sigma x s0) (l_e_st_esti sigma x t0)) s (l_imp_th4 (l_e_st_esti sigma s s0) (l_not (l_e_st_esti sigma s t0)) ses0 (l_weli (l_e_st_esti sigma s t0) set0)))))))). Time Defined. (* constant 419 *) Definition l_e_st_disj_th2 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (set0:l_e_st_esti sigma s t0), l_not (l_e_st_disj sigma t0 s0))))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (set0:l_e_st_esti sigma s t0) => l_e_st_disj_th1 sigma t0 s0 s set0 ses0)))))). Time Defined. (* constant 420 *) Definition l_e_st_issete1 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (i:l_e_is (l_e_st_set sigma) s0 t0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), l_e_st_esti sigma s t0)))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (i:l_e_is (l_e_st_set sigma) s0 t0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => l_e_isp (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_esti sigma s x) s0 t0 ses0 i)))))). Time Defined. (* constant 421 *) Definition l_e_st_issete2 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (i:l_e_is (l_e_st_set sigma) s0 t0), (forall (s:sigma), (forall (set0:l_e_st_esti sigma s t0), l_e_st_esti sigma s s0)))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (i:l_e_is (l_e_st_set sigma) s0 t0) => (fun (s:sigma) => (fun (set0:l_e_st_esti sigma s t0) => l_e_isp1 (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_esti sigma s x) t0 s0 set0 i)))))). Time Defined. (* constant 422 *) Definition l_e_st_isset_th1 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (i:l_e_is (l_e_st_set sigma) s0 t0), l_e_st_incl sigma s0 t0)))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (i:l_e_is (l_e_st_set sigma) s0 t0) => (fun (x:sigma) => (fun (y:l_e_st_esti sigma x s0) => l_e_st_issete1 sigma s0 t0 i x y)))))). Time Defined. (* constant 423 *) Definition l_e_st_isset_th2 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (i:l_e_is (l_e_st_set sigma) s0 t0), l_e_st_incl sigma t0 s0)))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (i:l_e_is (l_e_st_set sigma) s0 t0) => (fun (x:sigma) => (fun (y:l_e_st_esti sigma x t0) => l_e_st_issete2 sigma s0 t0 i x y)))))). Time Defined. (* constant 424 *) Axiom l_e_st_isseti : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (i:l_e_st_incl sigma s0 t0), (forall (j:l_e_st_incl sigma t0 s0), l_e_is (l_e_st_set sigma) s0 t0))))). (* constant 425 *) Definition l_e_st_isset_th3 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (n:l_not (l_e_st_esti sigma s t0)), l_not (l_e_is (l_e_st_set sigma) s0 t0))))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (n:l_not (l_e_st_esti sigma s t0)) => l_imp_th3 (l_e_is (l_e_st_set sigma) s0 t0) (l_e_st_esti sigma s t0) n (fun (t:l_e_is (l_e_st_set sigma) s0 t0) => l_e_st_issete1 sigma s0 t0 t s ses0))))))). Time Defined. (* constant 426 *) Definition l_e_st_isset_th4 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (n:l_not (l_e_st_esti sigma s t0)), l_not (l_e_is (l_e_st_set sigma) t0 s0))))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (n:l_not (l_e_st_esti sigma s t0)) => l_e_symnotis (l_e_st_set sigma) s0 t0 (l_e_st_isset_th3 sigma s0 t0 s ses0 n))))))). Time Defined. (* constant 427 *) Definition l_e_st_isset_nissetprop : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (s:sigma), Prop)))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (s:sigma) => l_and (l_e_st_esti sigma s s0) (l_not (l_e_st_esti sigma s t0)))))). Time Defined. (* constant 428 *) Definition l_e_st_isset_t1 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (s:sigma), (forall (n:l_not (l_e_st_isset_nissetprop sigma s0 t0 s)), (forall (e:l_e_st_esti sigma s s0), l_e_st_esti sigma s t0)))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (s:sigma) => (fun (n:l_not (l_e_st_isset_nissetprop sigma s0 t0 s)) => (fun (e:l_e_st_esti sigma s s0) => l_et (l_e_st_esti sigma s t0) (l_and_th3 (l_e_st_esti sigma s s0) (l_not (l_e_st_esti sigma s t0)) n e))))))). Time Defined. (* constant 429 *) Definition l_e_st_isset_t2 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)), (forall (m:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))), (forall (l:l_none sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)), (forall (s:sigma), l_not (l_e_st_isset_nissetprop sigma s0 t0 s)))))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)) => (fun (m:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))) => (fun (l:l_none sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)) => (fun (s:sigma) => l_some_th4 sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x) m s))))))). Time Defined. (* constant 430 *) Definition l_e_st_isset_t3 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)), (forall (m:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))), (forall (l:l_none sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)), (forall (s:sigma), l_not (l_e_st_isset_nissetprop sigma t0 s0 s)))))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)) => (fun (m:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))) => (fun (l:l_none sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)) => (fun (s:sigma) => l s))))))). Time Defined. (* constant 431 *) Definition l_e_st_isset_t4 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)), (forall (m:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))), (forall (l:l_none sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)), l_e_is (l_e_st_set sigma) s0 t0)))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)) => (fun (m:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))) => (fun (l:l_none sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)) => l_e_st_isseti sigma s0 t0 (fun (x:sigma) => (fun (y:l_e_st_esti sigma x s0) => l_e_st_isset_t1 sigma s0 t0 x (l_e_st_isset_t2 sigma s0 t0 n m l x) y)) (fun (x:sigma) => (fun (y:l_e_st_esti sigma x t0) => l_e_st_isset_t1 sigma t0 s0 x (l_e_st_isset_t3 sigma s0 t0 n m l x) y)))))))). Time Defined. (* constant 432 *) Definition l_e_st_isset_t5 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)), (forall (m:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))), l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)) => (fun (m:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))) => l_imp_th3 (l_none sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)) (l_e_is (l_e_st_set sigma) s0 t0) n (fun (y:l_none sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)) => l_e_st_isset_t4 sigma s0 t0 n m y)))))). Time Defined. (* constant 433 *) Definition l_e_st_isset_th5 : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), (forall (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)), l_or (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x)) (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)))))). exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (t0:l_e_st_set sigma) => (fun (n:l_not (l_e_is (l_e_st_set sigma) s0 t0)) => l_or_th1 (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x)) (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma t0 s0 x)) (fun (y:l_not (l_some sigma (fun (x:sigma) => l_e_st_isset_nissetprop sigma s0 t0 x))) => l_e_st_isset_t5 sigma s0 t0 n y))))). Time Defined. (* constant 434 *) Definition l_e_st_unmore : (forall (sigma:Type), (forall (alpha:Type), (forall (sa:(forall (x:alpha), l_e_st_set sigma)), l_e_st_set sigma))). exact (fun (sigma:Type) => (fun (alpha:Type) => (fun (sa:(forall (x:alpha), l_e_st_set sigma)) => l_e_st_setof sigma (fun (x:sigma) => l_some alpha (fun (y:alpha) => l_e_st_esti sigma x (sa y)))))). Time Defined. (* constant 435 *) Definition l_e_st_eunmore1 : (forall (sigma:Type), (forall (alpha:Type), (forall (sa:(forall (x:alpha), l_e_st_set sigma)), (forall (s:sigma), (forall (a:alpha), (forall (seasa:l_e_st_esti sigma s (sa a)), l_e_st_esti sigma s (l_e_st_unmore sigma alpha sa))))))). exact (fun (sigma:Type) => (fun (alpha:Type) => (fun (sa:(forall (x:alpha), l_e_st_set sigma)) => (fun (s:sigma) => (fun (a:alpha) => (fun (seasa:l_e_st_esti sigma s (sa a)) => l_e_st_estii sigma (fun (x:sigma) => l_some alpha (fun (y:alpha) => l_e_st_esti sigma x (sa y))) s (l_somei alpha (fun (y:alpha) => l_e_st_esti sigma s (sa y)) a seasa))))))). Time Defined. (* constant 436 *) Definition l_e_st_unmoreapp : (forall (sigma:Type), (forall (alpha:Type), (forall (sa:(forall (x:alpha), l_e_st_set sigma)), (forall (s:sigma), (forall (seun:l_e_st_esti sigma s (l_e_st_unmore sigma alpha sa)), (forall (x:Prop), (forall (x1:(forall (y:alpha), (forall (z:l_e_st_esti sigma s (sa y)), x))), x))))))). exact (fun (sigma:Type) => (fun (alpha:Type) => (fun (sa:(forall (x:alpha), l_e_st_set sigma)) => (fun (s:sigma) => (fun (seun:l_e_st_esti sigma s (l_e_st_unmore sigma alpha sa)) => (fun (x:Prop) => (fun (x1:(forall (y:alpha), (forall (z:l_e_st_esti sigma s (sa y)), x))) => l_someapp alpha (fun (y:alpha) => l_e_st_esti sigma s (sa y)) (l_e_st_estie sigma (fun (z:sigma) => l_some alpha (fun (y:alpha) => l_e_st_esti sigma z (sa y))) s seun) x x1))))))). Time Defined. (* constant 437 *) Definition l_e_st_eq_refr : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), r s s)))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => refr1 s)))))). Time Defined. (* constant 438 *) Definition l_e_st_eq_symr : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), r t s)))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => symr1 s t tsr)))))))). Time Defined. (* constant 439 *) Definition l_e_st_eq_trr : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (tsr:r s t), (forall (utr:r t u), r s u)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (tsr:r s t) => (fun (utr:r t u) => trr1 s t u tsr utr)))))))))). Time Defined. (* constant 440 *) Definition l_e_st_eq_tr1r : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (sur:r u s), (forall (tur:r u t), r s t)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (sur:r u s) => (fun (tur:r u t) => l_e_st_eq_trr sigma r refr1 symr1 trr1 s u t (l_e_st_eq_symr sigma r refr1 symr1 trr1 u s sur) tur)))))))))). Time Defined. (* constant 441 *) Definition l_e_st_eq_tr2r : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (u:sigma), (forall (usr:r s u), (forall (utr:r t u), r s t)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (u:sigma) => (fun (usr:r s u) => (fun (utr:r t u) => l_e_st_eq_trr sigma r refr1 symr1 trr1 s u t usr (l_e_st_eq_symr sigma r refr1 symr1 trr1 t u utr))))))))))). Time Defined. (* constant 442 *) Definition l_e_st_eq_ecelt : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), l_e_st_set sigma)))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => l_e_st_setof sigma (r s))))))). Time Defined. (* constant 443 *) Definition l_e_st_eq_1_th1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), l_e_st_esti sigma s (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => l_e_st_estii sigma (r s) s (l_e_st_eq_refr sigma r refr1 symr1 trr1 s))))))). Time Defined. (* constant 444 *) Definition l_e_st_eq_1_th2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), l_e_st_esti sigma t (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => l_e_st_estii sigma (r s) t tsr)))))))). Time Defined. (* constant 445 *) Definition l_e_st_eq_1_th3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (e:l_e_st_esti sigma t (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)), r s t)))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (e:l_e_st_esti sigma t (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)) => l_e_st_estie sigma (r s) t e)))))))). Time Defined. (* constant 446 *) Definition l_e_st_eq_1_t1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), (forall (u:sigma), (forall (e:l_e_st_esti sigma u (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)), l_e_st_esti sigma u (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => (fun (u:sigma) => (fun (e:l_e_st_esti sigma u (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)) => l_e_st_eq_1_th2 sigma r refr1 symr1 trr1 t u (l_e_st_eq_tr1r sigma r refr1 symr1 trr1 t u s tsr (l_e_st_eq_1_th3 sigma r refr1 symr1 trr1 s u e)))))))))))). Time Defined. (* constant 447 *) Definition l_e_st_eq_1_th4 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), l_e_is (l_e_st_set sigma) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => l_e_st_isseti sigma (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t) (fun (x:sigma) => (fun (y:l_e_st_esti sigma x (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)) => l_e_st_eq_1_t1 sigma r refr1 symr1 trr1 s t tsr x y)) (fun (x:sigma) => (fun (y:l_e_st_esti sigma x (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t)) => l_e_st_eq_1_t1 sigma r refr1 symr1 trr1 t s (l_e_st_eq_symr sigma r refr1 symr1 trr1 s t tsr) x y)))))))))). Time Defined. (* constant 448 *) Definition l_e_st_eq_1_t2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (n:l_not (r s t)), (forall (u:sigma), (forall (e:l_e_st_esti sigma u (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)), l_not (l_e_st_esti sigma u (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (n:l_not (r s t)) => (fun (u:sigma) => (fun (e:l_e_st_esti sigma u (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)) => l_imp_th3 (l_e_st_esti sigma u (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t)) (r s t) n (fun (x:l_e_st_esti sigma u (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t)) => l_e_st_eq_tr2r sigma r refr1 symr1 trr1 s t u (l_e_st_eq_1_th3 sigma r refr1 symr1 trr1 s u e) (l_e_st_eq_1_th3 sigma r refr1 symr1 trr1 t u x)))))))))))). Time Defined. (* constant 449 *) Definition l_e_st_eq_1_th5 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (n:l_not (r s t)), l_e_st_disj sigma (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (n:l_not (r s t)) => (fun (x:sigma) => (fun (y:l_e_st_esti sigma x (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)) => l_e_st_eq_1_t2 sigma r refr1 symr1 trr1 s t n x y)))))))))). Time Defined. (* constant 450 *) Definition l_e_st_eq_1_th6 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), l_e_st_nonempty sigma (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => l_e_st_nonemptyi sigma (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) s (l_e_st_eq_1_th1 sigma r refr1 symr1 trr1 s))))))). Time Defined. (* constant 451 *) Definition l_e_st_eq_ecp : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (s:sigma), Prop))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (s:sigma) => l_e_is (l_e_st_set sigma) s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)))))))). Time Defined. (* constant 452 *) Definition l_e_st_eq_anec : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), Prop)))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => l_some sigma (fun (x:sigma) => l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 x))))))). Time Defined. (* constant 453 *) Definition l_e_st_eq_2_th1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), l_e_st_eq_anec sigma r refr1 symr1 trr1 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => l_somei sigma (fun (x:sigma) => l_e_st_eq_ecp sigma r refr1 symr1 trr1 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) x) s (l_e_refis (l_e_st_set sigma) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)))))))). Time Defined. (* constant 454 *) Definition l_e_st_eq_2_t1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (t:sigma), (forall (e:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 t), l_e_st_esti sigma s (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (t:sigma) => (fun (e:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 t) => l_e_st_issete1 sigma s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t) e s ses0))))))))))). Time Defined. (* constant 455 *) Definition l_e_st_eq_2_t2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (t:sigma), (forall (e:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 t), l_e_is (l_e_st_set sigma) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (t:sigma) => (fun (e:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 t) => l_e_st_eq_1_th4 sigma r refr1 symr1 trr1 t s (l_e_st_eq_1_th3 sigma r refr1 symr1 trr1 t s (l_e_st_eq_2_t1 sigma r refr1 symr1 trr1 s0 ecs0 s ses0 t e))))))))))))). Time Defined. (* constant 456 *) Definition l_e_st_eq_2_t3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (t:sigma), (forall (e:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 t), l_e_is (l_e_st_set sigma) s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (t:sigma) => (fun (e:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 t) => l_e_tris (l_e_st_set sigma) s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) e (l_e_st_eq_2_t2 sigma r refr1 symr1 trr1 s0 ecs0 s ses0 t e)))))))))))). Time Defined. (* constant 457 *) Definition l_e_st_eq_2_th2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), l_e_is (l_e_st_set sigma) s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => l_someapp sigma (fun (x:sigma) => l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 x) ecs0 (l_e_is (l_e_st_set sigma) s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s)) (fun (x:sigma) => (fun (y:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 x) => l_e_st_eq_2_t3 sigma r refr1 symr1 trr1 s0 ecs0 s ses0 x y))))))))))). Time Defined. (* constant 458 *) Definition l_e_st_eq_2_th3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (t:sigma), (forall (tes0:l_e_st_esti sigma t s0), r s t))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (t:sigma) => (fun (tes0:l_e_st_esti sigma t s0) => l_e_st_eq_1_th3 sigma r refr1 symr1 trr1 s t (l_e_st_issete1 sigma s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_2_th2 sigma r refr1 symr1 trr1 s0 ecs0 s ses0) t tes0)))))))))))). Time Defined. (* constant 459 *) Definition l_e_st_eq_2_th4 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (t:sigma), (forall (tsr:r s t), l_e_st_esti sigma t s0))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (t:sigma) => (fun (tsr:r s t) => l_e_st_issete2 sigma s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_2_th2 sigma r refr1 symr1 trr1 s0 ecs0 s ses0) t (l_e_st_eq_1_th2 sigma r refr1 symr1 trr1 s t tsr)))))))))))). Time Defined. (* constant 460 *) Definition l_e_st_eq_2_t4 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (s:sigma), (forall (e:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 s), l_e_st_nonempty sigma s0))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (s:sigma) => (fun (e:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 s) => l_e_isp (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_nonempty sigma x) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) s0 (l_e_st_eq_1_th6 sigma r refr1 symr1 trr1 s) (l_e_symis (l_e_st_set sigma) s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) e)))))))))). Time Defined. (* constant 461 *) Definition l_e_st_eq_2_th5 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), l_e_st_nonempty sigma s0))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => l_someapp sigma (fun (x:sigma) => l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 x) ecs0 (l_e_st_nonempty sigma s0) (fun (x:sigma) => (fun (y:l_e_st_eq_ecp sigma r refr1 symr1 trr1 s0 x) => l_e_st_eq_2_t4 sigma r refr1 symr1 trr1 s0 ecs0 x y))))))))). Time Defined. (* constant 462 *) Definition l_e_st_eq_3_th1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (t0:l_e_st_set sigma), (forall (ect0:l_e_st_eq_anec sigma r refr1 symr1 trr1 t0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (t:sigma), (forall (tet0:l_e_st_esti sigma t t0), (forall (tsr:r s t), l_e_is (l_e_st_set sigma) s0 t0)))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (t0:l_e_st_set sigma) => (fun (ect0:l_e_st_eq_anec sigma r refr1 symr1 trr1 t0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (t:sigma) => (fun (tet0:l_e_st_esti sigma t t0) => (fun (tsr:r s t) => l_e_tr3is (l_e_st_set sigma) s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t) t0 (l_e_st_eq_2_th2 sigma r refr1 symr1 trr1 s0 ecs0 s ses0) (l_e_st_eq_1_th4 sigma r refr1 symr1 trr1 s t tsr) (l_e_symis (l_e_st_set sigma) t0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t) (l_e_st_eq_2_th2 sigma r refr1 symr1 trr1 t0 ect0 t tet0)))))))))))))))). Time Defined. (* constant 463 *) Definition l_e_st_eq_3_t1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (t0:l_e_st_set sigma), (forall (ect0:l_e_st_eq_anec sigma r refr1 symr1 trr1 t0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (t:sigma), (forall (tet0:l_e_st_esti sigma t t0), (forall (n:l_not (r s t)), l_e_st_disj sigma s0 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (t0:l_e_st_set sigma) => (fun (ect0:l_e_st_eq_anec sigma r refr1 symr1 trr1 t0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (t:sigma) => (fun (tet0:l_e_st_esti sigma t t0) => (fun (n:l_not (r s t)) => l_e_isp1 (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_disj sigma x (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t)) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) s0 (l_e_st_eq_1_th5 sigma r refr1 symr1 trr1 s t n) (l_e_st_eq_2_th2 sigma r refr1 symr1 trr1 s0 ecs0 s ses0))))))))))))))). Time Defined. (* constant 464 *) Definition l_e_st_eq_3_th2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (t0:l_e_st_set sigma), (forall (ect0:l_e_st_eq_anec sigma r refr1 symr1 trr1 t0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), (forall (t:sigma), (forall (tet0:l_e_st_esti sigma t t0), (forall (n:l_not (r s t)), l_e_st_disj sigma s0 t0)))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (t0:l_e_st_set sigma) => (fun (ect0:l_e_st_eq_anec sigma r refr1 symr1 trr1 t0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => (fun (t:sigma) => (fun (tet0:l_e_st_esti sigma t t0) => (fun (n:l_not (r s t)) => l_e_isp1 (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_disj sigma s0 x) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t) t0 (l_e_st_eq_3_t1 sigma r refr1 symr1 trr1 s0 ecs0 t0 ect0 s ses0 t tet0 n) (l_e_st_eq_2_th2 sigma r refr1 symr1 trr1 t0 ect0 t tet0))))))))))))))). Time Defined. (* constant 465 *) Definition l_e_st_eq_3_t2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (t0:l_e_st_set sigma), (forall (i:l_e_is (l_e_st_set sigma) s0 t0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), l_e_st_esti sigma s t0))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (t0:l_e_st_set sigma) => (fun (i:l_e_is (l_e_st_set sigma) s0 t0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => l_e_st_issete1 sigma s0 t0 i s ses0))))))))))). Time Defined. (* constant 466 *) Definition l_e_st_eq_3_t3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (t0:l_e_st_set sigma), (forall (i:l_e_is (l_e_st_set sigma) s0 t0), (forall (s:sigma), (forall (ses0:l_e_st_esti sigma s s0), l_not (l_e_st_disj sigma s0 t0)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (t0:l_e_st_set sigma) => (fun (i:l_e_is (l_e_st_set sigma) s0 t0) => (fun (s:sigma) => (fun (ses0:l_e_st_esti sigma s s0) => l_e_st_disj_th1 sigma s0 t0 s ses0 (l_e_st_eq_3_t2 sigma r refr1 symr1 trr1 s0 ecs0 t0 i s ses0)))))))))))). Time Defined. (* constant 467 *) Definition l_e_st_eq_3_th3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), (forall (t0:l_e_st_set sigma), (forall (i:l_e_is (l_e_st_set sigma) s0 t0), l_not (l_e_st_disj sigma s0 t0)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => (fun (t0:l_e_st_set sigma) => (fun (i:l_e_is (l_e_st_set sigma) s0 t0) => l_e_st_nonemptyapp sigma s0 (l_e_st_eq_2_th5 sigma r refr1 symr1 trr1 s0 ecs0) (l_not (l_e_st_disj sigma s0 t0)) (fun (x:sigma) => (fun (y:l_e_st_esti sigma x s0) => l_e_st_eq_3_t3 sigma r refr1 symr1 trr1 s0 ecs0 t0 i x y))))))))))). Time Defined. (* constant 468 *) Definition l_e_st_eq_ect : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), Type))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => l_e_ot (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_eq_anec sigma r refr1 symr1 trr1 x)))))). Time Defined. (* constant 469 *) Definition l_e_st_eq_ectset : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s0:l_e_st_set sigma), (forall (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0), l_e_st_eq_ect sigma r refr1 symr1 trr1))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s0:l_e_st_set sigma) => (fun (ecs0:l_e_st_eq_anec sigma r refr1 symr1 trr1 s0) => l_e_out (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_eq_anec sigma r refr1 symr1 trr1 x) s0 ecs0))))))). Time Defined. (* constant 470 *) Definition l_e_st_eq_ectelt : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), l_e_st_eq_ect sigma r refr1 symr1 trr1)))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => l_e_st_eq_ectset sigma r refr1 symr1 trr1 (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_2_th1 sigma r refr1 symr1 trr1 s))))))). Time Defined. (* constant 471 *) Definition l_e_st_eq_ecect : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), l_e_st_set sigma)))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_in (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_eq_anec sigma r refr1 symr1 trr1 x) e)))))). Time Defined. (* constant 472 *) Definition l_e_st_eq_4_th1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), l_e_st_eq_anec sigma r refr1 symr1 trr1 (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_inp (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_eq_anec sigma r refr1 symr1 trr1 x) e)))))). Time Defined. (* constant 473 *) Definition l_e_st_eq_4_th2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), l_e_st_nonempty sigma (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_st_eq_2_th5 sigma r refr1 symr1 trr1 (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_4_th1 sigma r refr1 symr1 trr1 e))))))). Time Defined. (* constant 474 *) Definition l_e_st_eq_4_th3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (x:Prop), (forall (x1:(forall (y:sigma), (forall (z:l_e_st_esti sigma y (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), x))), x)))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (x:Prop) => (fun (x1:(forall (y:sigma), (forall (z:l_e_st_esti sigma y (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), x))) => l_e_st_nonemptyapp sigma (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_4_th2 sigma r refr1 symr1 trr1 e) x x1)))))))). Time Defined. (* constant 475 *) Definition l_e_st_eq_4_th4 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), l_e_is (l_e_st_set sigma) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s)))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => l_e_isinout (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_eq_anec sigma r refr1 symr1 trr1 x) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_2_th1 sigma r refr1 symr1 trr1 s))))))). Time Defined. (* constant 476 *) Definition l_e_st_eq_4_th5 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s)))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => l_e_st_issete1 sigma (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s)) (l_e_st_eq_4_th4 sigma r refr1 symr1 trr1 s) s (l_e_st_eq_1_th1 sigma r refr1 symr1 trr1 s))))))). Time Defined. (* constant 477 *) Definition l_e_st_eq_4_th6 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), l_e_st_esti sigma s (l_e_st_unmore sigma (l_e_st_eq_ect sigma r refr1 symr1 trr1) (fun (x:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_st_eq_ecect sigma r refr1 symr1 trr1 x)))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => l_e_st_eunmore1 sigma (l_e_st_eq_ect sigma r refr1 symr1 trr1) (fun (x:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_st_eq_ecect sigma r refr1 symr1 trr1 x) s (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_4_th5 sigma r refr1 symr1 trr1 s))))))). Time Defined. (* constant 478 *) Definition l_e_st_eq_4_th7 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (t:sigma), (forall (tee:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), r s t)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (t:sigma) => (fun (tee:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_e_st_eq_2_th3 sigma r refr1 symr1 trr1 (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_4_th1 sigma r refr1 symr1 trr1 e) s see t tee)))))))))). Time Defined. (* constant 479 *) Definition l_e_st_eq_4_th8 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (t:sigma), (forall (tsr:r s t), l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (t:sigma) => (fun (tsr:r s t) => l_e_st_eq_2_th4 sigma r refr1 symr1 trr1 (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_4_th1 sigma r refr1 symr1 trr1 e) s see t tsr)))))))))). Time Defined. (* constant 480 *) Definition l_e_st_eq_5_th1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (i:l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) e f), l_e_is (l_e_st_set sigma) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (i:l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) e f) => l_e_isini (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_eq_anec sigma r refr1 symr1 trr1 x) e f i)))))))). Time Defined. (* constant 481 *) Definition l_e_st_eq_5_th2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (i:l_e_is (l_e_st_set sigma) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)), l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) e f)))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (i:l_e_is (l_e_st_set sigma) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)) => l_e_isine (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_eq_anec sigma r refr1 symr1 trr1 x) e f i)))))))). Time Defined. (* constant 482 *) Definition l_e_st_eq_5_th3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (t:sigma), (forall (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)), (forall (tsr:r s t), l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) e f)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (t:sigma) => (fun (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)) => (fun (tsr:r s t) => l_e_st_eq_5_th2 sigma r refr1 symr1 trr1 e f (l_e_st_eq_3_th1 sigma r refr1 symr1 trr1 (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_4_th1 sigma r refr1 symr1 trr1 e) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f) (l_e_st_eq_4_th1 sigma r refr1 symr1 trr1 f) s see t tef tsr))))))))))))). Time Defined. (* constant 483 *) Definition l_e_st_eq_5_th4 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (i:l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) e f), l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (i:l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) e f) => l_e_st_issete1 sigma (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e) (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f) (l_e_st_eq_5_th1 sigma r refr1 symr1 trr1 e f i) s see)))))))))). Time Defined. (* constant 484 *) Definition l_e_st_eq_5_th5 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (t:sigma), (forall (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)), (forall (i:l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) e f), r s t)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (t:sigma) => (fun (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)) => (fun (i:l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) e f) => l_e_st_eq_2_th3 sigma r refr1 symr1 trr1 (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f) (l_e_st_eq_4_th1 sigma r refr1 symr1 trr1 f) s (l_e_st_eq_5_th4 sigma r refr1 symr1 trr1 e f s see i) t tef)))))))))))). Time Defined. (* constant 485 *) Definition l_e_st_eq_5_th6 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), l_e_is (l_e_st_eq_ect sigma r refr1 symr1 trr1) (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 t))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => l_e_isouti (l_e_st_set sigma) (fun (x:l_e_st_set sigma) => l_e_st_eq_anec sigma r refr1 symr1 trr1 x) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_2_th1 sigma r refr1 symr1 trr1 s) (l_e_st_eq_ecelt sigma r refr1 symr1 trr1 t) (l_e_st_eq_2_th1 sigma r refr1 symr1 trr1 t) (l_e_st_eq_1_th4 sigma r refr1 symr1 trr1 s t tsr))))))))). Time Defined. (* constant 486 *) Definition l_e_st_eq_fixfu : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), Prop))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (forall (x:sigma), (forall (y:sigma), (forall (z:r x y), l_e_is alpha (fu x) (fu y))))))))))). Time Defined. (* constant 487 *) Definition l_e_st_eq_10_prop1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (s:sigma), Prop))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (s:sigma) => l_and (l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) (l_e_is alpha (fu s) a1)))))))))))). Time Defined. (* constant 488 *) Definition l_e_st_eq_10_prop2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), Prop)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => l_some sigma (fun (x:sigma) => l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 x))))))))))). Time Defined. (* constant 489 *) Definition l_e_st_eq_10_t1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e (fu s) s))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_andi (l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) (l_e_is alpha (fu s) (fu s)) see (l_e_refis alpha (fu s))))))))))))). Time Defined. (* constant 490 *) Definition l_e_st_eq_10_t2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e (fu s)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_somei sigma (fun (x:sigma) => l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e (fu s) x) s (l_e_st_eq_10_t1 sigma r refr1 symr1 trr1 alpha fu ff e s see)))))))))))). Time Defined. (* constant 491 *) Definition l_e_st_eq_10_t3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_some alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_somei alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x) (fu s) (l_e_st_eq_10_t2 sigma r refr1 symr1 trr1 alpha fu ff e s see)))))))))))). Time Defined. (* constant 492 *) Definition l_e_st_eq_10_t4 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), l_some alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_st_eq_4_th3 sigma r refr1 symr1 trr1 e (l_some alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x)) (fun (x:sigma) => (fun (y:l_e_st_esti sigma x (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_e_st_eq_10_t3 sigma r refr1 symr1 trr1 alpha fu ff e x y))))))))))). Time Defined. (* constant 493 *) Definition l_e_st_eq_10_t5 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (b1:alpha), (forall (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1), (forall (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1), (forall (s:sigma), (forall (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s), (forall (t:sigma), (forall (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t), l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)))))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (b1:alpha) => (fun (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1) => (fun (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1) => (fun (s:sigma) => (fun (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s) => (fun (t:sigma) => (fun (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t) => l_ande1 (l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) (l_e_is alpha (fu s) a1) pa1s))))))))))))))))). Time Defined. (* constant 494 *) Definition l_e_st_eq_10_t6 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (b1:alpha), (forall (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1), (forall (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1), (forall (s:sigma), (forall (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s), (forall (t:sigma), (forall (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t), l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)))))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (b1:alpha) => (fun (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1) => (fun (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1) => (fun (s:sigma) => (fun (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s) => (fun (t:sigma) => (fun (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t) => l_ande1 (l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) (l_e_is alpha (fu t) b1) pb1t))))))))))))))))). Time Defined. (* constant 495 *) Definition l_e_st_eq_10_t7 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (b1:alpha), (forall (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1), (forall (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1), (forall (s:sigma), (forall (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s), (forall (t:sigma), (forall (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t), r s t))))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (b1:alpha) => (fun (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1) => (fun (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1) => (fun (s:sigma) => (fun (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s) => (fun (t:sigma) => (fun (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t) => l_e_st_eq_4_th7 sigma r refr1 symr1 trr1 e s (l_e_st_eq_10_t5 sigma r refr1 symr1 trr1 alpha fu ff e a1 b1 pa1 pb1 s pa1s t pb1t) t (l_e_st_eq_10_t6 sigma r refr1 symr1 trr1 alpha fu ff e a1 b1 pa1 pb1 s pa1s t pb1t)))))))))))))))))). Time Defined. (* constant 496 *) Definition l_e_st_eq_10_t8 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (b1:alpha), (forall (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1), (forall (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1), (forall (s:sigma), (forall (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s), (forall (t:sigma), (forall (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t), l_e_is alpha (fu s) a1))))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (b1:alpha) => (fun (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1) => (fun (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1) => (fun (s:sigma) => (fun (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s) => (fun (t:sigma) => (fun (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t) => l_ande2 (l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) (l_e_is alpha (fu s) a1) pa1s))))))))))))))))). Time Defined. (* constant 497 *) Definition l_e_st_eq_10_t9 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (b1:alpha), (forall (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1), (forall (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1), (forall (s:sigma), (forall (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s), (forall (t:sigma), (forall (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t), l_e_is alpha (fu t) b1))))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (b1:alpha) => (fun (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1) => (fun (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1) => (fun (s:sigma) => (fun (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s) => (fun (t:sigma) => (fun (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t) => l_ande2 (l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) (l_e_is alpha (fu t) b1) pb1t))))))))))))))))). Time Defined. (* constant 498 *) Definition l_e_st_eq_10_t10 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (b1:alpha), (forall (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1), (forall (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1), (forall (s:sigma), (forall (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s), (forall (t:sigma), (forall (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t), l_e_is alpha a1 b1))))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (b1:alpha) => (fun (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1) => (fun (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1) => (fun (s:sigma) => (fun (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s) => (fun (t:sigma) => (fun (pb1t:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 t) => l_e_tr3is alpha a1 (fu s) (fu t) b1 (l_e_symis alpha (fu s) a1 (l_e_st_eq_10_t8 sigma r refr1 symr1 trr1 alpha fu ff e a1 b1 pa1 pb1 s pa1s t pb1t)) (ff s t (l_e_st_eq_10_t7 sigma r refr1 symr1 trr1 alpha fu ff e a1 b1 pa1 pb1 s pa1s t pb1t)) (l_e_st_eq_10_t9 sigma r refr1 symr1 trr1 alpha fu ff e a1 b1 pa1 pb1 s pa1s t pb1t)))))))))))))))))). Time Defined. (* constant 499 *) Definition l_e_st_eq_10_t11 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (b1:alpha), (forall (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1), (forall (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1), (forall (s:sigma), (forall (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s), l_e_is alpha a1 b1))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (b1:alpha) => (fun (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1) => (fun (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1) => (fun (s:sigma) => (fun (pa1s:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 s) => l_someapp sigma (fun (x:sigma) => l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 x) pb1 (l_e_is alpha a1 b1) (fun (x:sigma) => (fun (y:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e b1 x) => l_e_st_eq_10_t10 sigma r refr1 symr1 trr1 alpha fu ff e a1 b1 pa1 pb1 s pa1s x y))))))))))))))))). Time Defined. (* constant 500 *) Definition l_e_st_eq_10_t12 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (a1:alpha), (forall (b1:alpha), (forall (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1), (forall (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1), l_e_is alpha a1 b1))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (a1:alpha) => (fun (b1:alpha) => (fun (pa1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e a1) => (fun (pb1:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e b1) => l_someapp sigma (fun (x:sigma) => l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 x) pa1 (l_e_is alpha a1 b1) (fun (x:sigma) => (fun (y:l_e_st_eq_10_prop1 sigma r refr1 symr1 trr1 alpha fu ff e a1 x) => l_e_st_eq_10_t11 sigma r refr1 symr1 trr1 alpha fu ff e a1 b1 pa1 pb1 x y))))))))))))))). Time Defined. (* constant 501 *) Definition l_e_st_eq_10_t13 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), l_e_amone alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (x:alpha) => (fun (y:alpha) => (fun (u:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x) => (fun (v:l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e y) => l_e_st_eq_10_t12 sigma r refr1 symr1 trr1 alpha fu ff e x y u v))))))))))))). Time Defined. (* constant 502 *) Definition l_e_st_eq_10_t14 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), l_e_one alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_onei alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x) (l_e_st_eq_10_t13 sigma r refr1 symr1 trr1 alpha fu ff e) (l_e_st_eq_10_t4 sigma r refr1 symr1 trr1 alpha fu ff e)))))))))). Time Defined. (* constant 503 *) Definition l_e_st_eq_indeq : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), alpha))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_ind alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x) (l_e_st_eq_10_t14 sigma r refr1 symr1 trr1 alpha fu ff e)))))))))). Time Defined. (* constant 504 *) Definition l_e_st_eq_10_th1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), l_some sigma (fun (x:sigma) => l_and (l_e_st_esti sigma x (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) (l_e_is alpha (fu x) (l_e_st_eq_indeq sigma r refr1 symr1 trr1 alpha fu ff e)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_oneax alpha (fun (x:alpha) => l_e_st_eq_10_prop2 sigma r refr1 symr1 trr1 alpha fu ff e x) (l_e_st_eq_10_t14 sigma r refr1 symr1 trr1 alpha fu ff e)))))))))). Time Defined. (* constant 505 *) Definition l_e_st_eq_10_th2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_e_is alpha (fu s) (l_e_st_eq_indeq sigma r refr1 symr1 trr1 alpha fu ff e)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_e_st_eq_10_t12 sigma r refr1 symr1 trr1 alpha fu ff e (fu s) (l_e_st_eq_indeq sigma r refr1 symr1 trr1 alpha fu ff e) (l_e_st_eq_10_t2 sigma r refr1 symr1 trr1 alpha fu ff e s see) (l_e_st_eq_10_th1 sigma r refr1 symr1 trr1 alpha fu ff e)))))))))))). Time Defined. (* constant 506 *) Definition l_e_st_eq_10_th3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu:(forall (x:sigma), alpha)), (forall (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu), (forall (s:sigma), l_e_is alpha (fu s) (l_e_st_eq_indeq sigma r refr1 symr1 trr1 alpha fu ff (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu:(forall (x:sigma), alpha)) => (fun (ff:l_e_st_eq_fixfu sigma r refr1 symr1 trr1 alpha fu) => (fun (s:sigma) => l_e_st_eq_10_th2 sigma r refr1 symr1 trr1 alpha fu ff (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s) s (l_e_st_eq_4_th5 sigma r refr1 symr1 trr1 s)))))))))). Time Defined. (* constant 507 *) Definition l_e_st_eq_fixfu2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), Prop))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (u:sigma), (forall (v:r x y), (forall (w:r z u), l_e_is alpha (fu2 x z) (fu2 y u)))))))))))))). Time Defined. (* constant 508 *) Definition l_e_st_eq_11_t1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), (forall (u:sigma), l_e_is alpha (fu2 s u) (fu2 t u))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => (fun (u:sigma) => ff2 s t u u tsr (l_e_st_eq_refr sigma r refr1 symr1 trr1 u))))))))))))). Time Defined. (* constant 509 *) Definition l_e_st_eq_11_t2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), l_e_is (forall (x:sigma), alpha) (fu2 s) (fu2 t)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => l_e_fisi sigma alpha (fu2 s) (fu2 t) (fun (x:sigma) => l_e_st_eq_11_t1 sigma r refr1 symr1 trr1 alpha fu2 ff2 s t tsr x)))))))))))). Time Defined. (* constant 510 *) Definition l_e_st_eq_11_i : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (x:sigma), alpha)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_st_eq_indeq sigma r refr1 symr1 trr1 (forall (x:sigma), alpha) fu2 (fun (x:sigma) => (fun (y:sigma) => (fun (z:r x y) => l_e_st_eq_11_t2 sigma r refr1 symr1 trr1 alpha fu2 ff2 x y z))) e))))))))). Time Defined. (* constant 511 *) Definition l_e_st_eq_11_t3 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), (forall (u:sigma), (forall (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_e_is (forall (x:sigma), alpha) (fu2 u) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => (fun (u:sigma) => (fun (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_e_st_eq_10_th2 sigma r refr1 symr1 trr1 (forall (x:sigma), alpha) fu2 (fun (x:sigma) => (fun (y:sigma) => (fun (z:r x y) => l_e_st_eq_11_t2 sigma r refr1 symr1 trr1 alpha fu2 ff2 x y z))) e u uee)))))))))))))). Time Defined. (* constant 512 *) Definition l_e_st_eq_11_t4 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), (forall (u:sigma), (forall (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_e_is alpha (fu2 u s) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e s))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => (fun (u:sigma) => (fun (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_e_fise sigma alpha (fu2 u) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e) (l_e_st_eq_11_t3 sigma r refr1 symr1 trr1 alpha fu2 ff2 e s t tsr u uee) s)))))))))))))). Time Defined. (* constant 513 *) Definition l_e_st_eq_11_t5 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), (forall (u:sigma), (forall (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_e_is alpha (fu2 u t) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e t))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => (fun (u:sigma) => (fun (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_e_fise sigma alpha (fu2 u) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e) (l_e_st_eq_11_t3 sigma r refr1 symr1 trr1 alpha fu2 ff2 e s t tsr u uee) t)))))))))))))). Time Defined. (* constant 514 *) Definition l_e_st_eq_11_t6 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), (forall (u:sigma), (forall (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_e_is alpha (fu2 u s) (fu2 u t))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => (fun (u:sigma) => (fun (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => ff2 u u s t (l_e_st_eq_refr sigma r refr1 symr1 trr1 u) tsr)))))))))))))). Time Defined. (* constant 515 *) Definition l_e_st_eq_11_t7 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), (forall (u:sigma), (forall (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), l_e_is alpha (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e s) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e t))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => (fun (u:sigma) => (fun (uee:l_e_st_esti sigma u (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_e_tr3is alpha (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e s) (fu2 u s) (fu2 u t) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e t) (l_e_symis alpha (fu2 u s) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e s) (l_e_st_eq_11_t4 sigma r refr1 symr1 trr1 alpha fu2 ff2 e s t tsr u uee)) (l_e_st_eq_11_t6 sigma r refr1 symr1 trr1 alpha fu2 ff2 e s t tsr u uee) (l_e_st_eq_11_t5 sigma r refr1 symr1 trr1 alpha fu2 ff2 e s t tsr u uee))))))))))))))). Time Defined. (* constant 516 *) Definition l_e_st_eq_11_t8 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (t:sigma), (forall (tsr:r s t), l_e_is alpha (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e s) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e t))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (t:sigma) => (fun (tsr:r s t) => l_e_st_eq_4_th3 sigma r refr1 symr1 trr1 e (l_e_is alpha (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e s) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e t)) (fun (x:sigma) => (fun (y:l_e_st_esti sigma x (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => l_e_st_eq_11_t7 sigma r refr1 symr1 trr1 alpha fu2 ff2 e s t tsr x y)))))))))))))). Time Defined. (* constant 517 *) Definition l_e_st_eq_indeq2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), alpha)))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => l_e_st_eq_indeq sigma r refr1 symr1 trr1 alpha (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e) (fun (x:sigma) => (fun (y:sigma) => (fun (z:r x y) => l_e_st_eq_11_t8 sigma r refr1 symr1 trr1 alpha fu2 ff2 e x y z))) f)))))))))). Time Defined. (* constant 518 *) Definition l_e_st_eq_11_t9 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (t:sigma), (forall (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)), l_e_is alpha (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e t) (l_e_st_eq_indeq2 sigma r refr1 symr1 trr1 alpha fu2 ff2 e f))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (t:sigma) => (fun (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)) => l_e_st_eq_10_th2 sigma r refr1 symr1 trr1 alpha (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e) (fun (x:sigma) => (fun (y:sigma) => (fun (z:r x y) => l_e_st_eq_11_t8 sigma r refr1 symr1 trr1 alpha fu2 ff2 e x y z))) f t tef)))))))))))))). Time Defined. (* constant 519 *) Definition l_e_st_eq_11_t10 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (t:sigma), (forall (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)), l_e_is (forall (x:sigma), alpha) (fu2 s) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (t:sigma) => (fun (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)) => l_e_st_eq_10_th2 sigma r refr1 symr1 trr1 (forall (x:sigma), alpha) fu2 (fun (x:sigma) => (fun (y:sigma) => (fun (z:r x y) => l_e_st_eq_11_t2 sigma r refr1 symr1 trr1 alpha fu2 ff2 x y z))) e s see)))))))))))))). Time Defined. (* constant 520 *) Definition l_e_st_eq_11_t11 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (t:sigma), (forall (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)), l_e_is alpha (fu2 s t) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e t))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (t:sigma) => (fun (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)) => l_e_fise sigma alpha (fu2 s) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e) (l_e_st_eq_11_t10 sigma r refr1 symr1 trr1 alpha fu2 ff2 e f s see t tef) t)))))))))))))). Time Defined. (* constant 521 *) Definition l_e_st_eq_11_th1 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (e:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (f:l_e_st_eq_ect sigma r refr1 symr1 trr1), (forall (s:sigma), (forall (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)), (forall (t:sigma), (forall (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)), l_e_is alpha (fu2 s t) (l_e_st_eq_indeq2 sigma r refr1 symr1 trr1 alpha fu2 ff2 e f))))))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (e:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (f:l_e_st_eq_ect sigma r refr1 symr1 trr1) => (fun (s:sigma) => (fun (see:l_e_st_esti sigma s (l_e_st_eq_ecect sigma r refr1 symr1 trr1 e)) => (fun (t:sigma) => (fun (tef:l_e_st_esti sigma t (l_e_st_eq_ecect sigma r refr1 symr1 trr1 f)) => l_e_tris alpha (fu2 s t) (l_e_st_eq_11_i sigma r refr1 symr1 trr1 alpha fu2 ff2 e t) (l_e_st_eq_indeq2 sigma r refr1 symr1 trr1 alpha fu2 ff2 e f) (l_e_st_eq_11_t11 sigma r refr1 symr1 trr1 alpha fu2 ff2 e f s see t tef) (l_e_st_eq_11_t9 sigma r refr1 symr1 trr1 alpha fu2 ff2 e f s see t tef))))))))))))))). Time Defined. (* constant 522 *) Definition l_e_st_eq_11_th2 : (forall (sigma:Type), (forall (r:(forall (x:sigma), (forall (y:sigma), Prop))), (forall (refr1:(forall (x:sigma), r x x)), (forall (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))), (forall (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))), (forall (alpha:Type), (forall (fu2:(forall (x:sigma), (forall (y:sigma), alpha))), (forall (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2), (forall (s:sigma), (forall (t:sigma), l_e_is alpha (fu2 s t) (l_e_st_eq_indeq2 sigma r refr1 symr1 trr1 alpha fu2 ff2 (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 t)))))))))))). exact (fun (sigma:Type) => (fun (r:(forall (x:sigma), (forall (y:sigma), Prop))) => (fun (refr1:(forall (x:sigma), r x x)) => (fun (symr1:(forall (x:sigma), (forall (y:sigma), (forall (t:r x y), r y x)))) => (fun (trr1:(forall (x:sigma), (forall (y:sigma), (forall (z:sigma), (forall (t:r x y), (forall (u:r y z), r x z)))))) => (fun (alpha:Type) => (fun (fu2:(forall (x:sigma), (forall (y:sigma), alpha))) => (fun (ff2:l_e_st_eq_fixfu2 sigma r refr1 symr1 trr1 alpha fu2) => (fun (s:sigma) => (fun (t:sigma) => l_e_st_eq_11_th1 sigma r refr1 symr1 trr1 alpha fu2 ff2 (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 s) (l_e_st_eq_ectelt sigma r refr1 symr1 trr1 t) s (l_e_st_eq_4_th5 sigma r refr1 symr1 trr1 s) t (l_e_st_eq_4_th5 sigma r refr1 symr1 trr1 t))))))))))). Time Defined. (* constant 523 *) Axiom l_e_st_eq_landau_n_nat : Type. (* constant 524 *) Definition l_e_st_eq_landau_n_is : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_is l_e_st_eq_landau_n_nat x y)). Time Defined. (* constant 525 *) Definition l_e_st_eq_landau_n_nis : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_not (l_e_st_eq_landau_n_is x y))). Time Defined. (* constant 526 *) Definition l_e_st_eq_landau_n_in : (forall (x:l_e_st_eq_landau_n_nat), (forall (s:l_e_st_set l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (s:l_e_st_set l_e_st_eq_landau_n_nat) => l_e_st_esti l_e_st_eq_landau_n_nat x s)). Time Defined. (* constant 527 *) Definition l_e_st_eq_landau_n_some : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => l_some l_e_st_eq_landau_n_nat p). Time Defined. (* constant 528 *) Definition l_e_st_eq_landau_n_all : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => l_all l_e_st_eq_landau_n_nat p). Time Defined. (* constant 529 *) Definition l_e_st_eq_landau_n_one : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => l_e_one l_e_st_eq_landau_n_nat p). Time Defined. (* constant 530 *) Axiom l_e_st_eq_landau_n_1 : l_e_st_eq_landau_n_nat. (* constant 531 *) Axiom l_e_st_eq_landau_n_suc : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat). (* constant 532 *) Definition l_e_st_eq_landau_n_ax2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_suc x y i))). Time Defined. (* constant 533 *) Axiom l_e_st_eq_landau_n_ax3 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_suc x) l_e_st_eq_landau_n_1). (* constant 534 *) Axiom l_e_st_eq_landau_n_ax4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc y)), l_e_st_eq_landau_n_is x y))). (* constant 535 *) Definition l_e_st_eq_landau_n_cond1 : (forall (s:l_e_st_set l_e_st_eq_landau_n_nat), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_in l_e_st_eq_landau_n_1 s). Time Defined. (* constant 536 *) Definition l_e_st_eq_landau_n_cond2 : (forall (s:l_e_st_set l_e_st_eq_landau_n_nat), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_all (fun (x:l_e_st_eq_landau_n_nat) => l_imp (l_e_st_eq_landau_n_in x s) (l_e_st_eq_landau_n_in (l_e_st_eq_landau_n_suc x) s))). Time Defined. (* constant 537 *) Axiom l_e_st_eq_landau_n_ax5 : (forall (s:l_e_st_set l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_cond1 s), (forall (v:l_e_st_eq_landau_n_cond2 s), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_in x s)))). (* constant 538 *) Definition l_e_st_eq_landau_n_i1_s : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n1p:p l_e_st_eq_landau_n_1), (forall (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_set l_e_st_eq_landau_n_nat)))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n1p:p l_e_st_eq_landau_n_1) => (fun (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_setof l_e_st_eq_landau_n_nat p)))). Time Defined. (* constant 539 *) Definition l_e_st_eq_landau_n_i1_t1 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n1p:p l_e_st_eq_landau_n_1), (forall (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_cond1 (l_e_st_eq_landau_n_i1_s p n1p xsp x))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n1p:p l_e_st_eq_landau_n_1) => (fun (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_estii l_e_st_eq_landau_n_nat p l_e_st_eq_landau_n_1 n1p)))). Time Defined. (* constant 540 *) Definition l_e_st_eq_landau_n_i1_t2 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n1p:p l_e_st_eq_landau_n_1), (forall (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (yes:l_e_st_eq_landau_n_in y (l_e_st_eq_landau_n_i1_s p n1p xsp x)), p y)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n1p:p l_e_st_eq_landau_n_1) => (fun (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (yes:l_e_st_eq_landau_n_in y (l_e_st_eq_landau_n_i1_s p n1p xsp x)) => l_e_st_estie l_e_st_eq_landau_n_nat p y yes)))))). Time Defined. (* constant 541 *) Definition l_e_st_eq_landau_n_i1_t3 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n1p:p l_e_st_eq_landau_n_1), (forall (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (yes:l_e_st_eq_landau_n_in y (l_e_st_eq_landau_n_i1_s p n1p xsp x)), l_e_st_eq_landau_n_in (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_i1_s p n1p xsp x))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n1p:p l_e_st_eq_landau_n_1) => (fun (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (yes:l_e_st_eq_landau_n_in y (l_e_st_eq_landau_n_i1_s p n1p xsp x)) => l_e_st_estii l_e_st_eq_landau_n_nat p (l_e_st_eq_landau_n_suc y) (xsp y (l_e_st_eq_landau_n_i1_t2 p n1p xsp x y yes)))))))). Time Defined. (* constant 542 *) Definition l_e_st_eq_landau_n_i1_t4 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n1p:p l_e_st_eq_landau_n_1), (forall (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_in x (l_e_st_eq_landau_n_i1_s p n1p xsp x))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n1p:p l_e_st_eq_landau_n_1) => (fun (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ax5 (l_e_st_eq_landau_n_i1_s p n1p xsp x) (l_e_st_eq_landau_n_i1_t1 p n1p xsp x) (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_in y (l_e_st_eq_landau_n_i1_s p n1p xsp x)) => l_e_st_eq_landau_n_i1_t3 p n1p xsp x y u)) x)))). Time Defined. (* constant 543 *) Definition l_e_st_eq_landau_n_induction : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n1p:p l_e_st_eq_landau_n_1), (forall (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))), (forall (x:l_e_st_eq_landau_n_nat), p x)))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n1p:p l_e_st_eq_landau_n_1) => (fun (xsp:(forall (x:l_e_st_eq_landau_n_nat), (forall (y:p x), p (l_e_st_eq_landau_n_suc x)))) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_estie l_e_st_eq_landau_n_nat p x (l_e_st_eq_landau_n_i1_t4 p n1p xsp x))))). Time Defined. (* constant 544 *) Definition l_e_st_eq_landau_n_21_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x y), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc y)), l_e_st_eq_landau_n_is x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x y) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc y)) => l_e_st_eq_landau_n_ax4 x y i)))). Time Defined. (* constant 545 *) Definition l_e_st_eq_landau_n_satz1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x y), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x y) => l_imp_th3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_is x y) n (fun (u:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc y)) => l_e_st_eq_landau_n_21_t1 x y n u)))). Time Defined. (* constant 546 *) Definition l_e_st_eq_landau_n_22_prop1 : (forall (x:l_e_st_eq_landau_n_nat), Prop). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_suc x) x). Time Defined. (* constant 547 *) Definition l_e_st_eq_landau_n_22_t1 : l_e_st_eq_landau_n_22_prop1 l_e_st_eq_landau_n_1. exact (l_e_st_eq_landau_n_ax3 l_e_st_eq_landau_n_1). Time Defined. (* constant 548 *) Definition l_e_st_eq_landau_n_22_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_22_prop1 x), l_e_st_eq_landau_n_22_prop1 (l_e_st_eq_landau_n_suc x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_22_prop1 x) => l_e_st_eq_landau_n_satz1 (l_e_st_eq_landau_n_suc x) x p)). Time Defined. (* constant 549 *) Definition l_e_st_eq_landau_n_satz2 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_suc x) x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_22_prop1 y) l_e_st_eq_landau_n_22_t1 (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_22_prop1 y) => l_e_st_eq_landau_n_22_t2 y u)) x). Time Defined. (* constant 550 *) Definition l_e_st_eq_landau_n_23_prop1 : (forall (x:l_e_st_eq_landau_n_nat), Prop). exact (fun (x:l_e_st_eq_landau_n_nat) => l_or (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)))). Time Defined. (* constant 551 *) Definition l_e_st_eq_landau_n_23_t1 : l_e_st_eq_landau_n_23_prop1 l_e_st_eq_landau_n_1. exact (l_ori1 (l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_suc u))) (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1)). Time Defined. (* constant 552 *) Definition l_e_st_eq_landau_n_23_t2 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc u))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_somei l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc u)) x (l_e_refis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc x))). Time Defined. (* constant 553 *) Definition l_e_st_eq_landau_n_23_t3 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_23_prop1 (l_e_st_eq_landau_n_suc x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_ori2 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_suc u))) (l_e_st_eq_landau_n_23_t2 x)). Time Defined. (* constant 554 *) Definition l_e_st_eq_landau_n_23_t4 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_23_prop1 x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_23_prop1 y) l_e_st_eq_landau_n_23_t1 (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_23_prop1 y) => l_e_st_eq_landau_n_23_t3 y)) x). Time Defined. (* constant 555 *) Definition l_e_st_eq_landau_n_satz3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_ore2 (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u))) (l_e_st_eq_landau_n_23_t4 x) n)). Time Defined. (* constant 556 *) Definition l_e_st_eq_landau_n_23_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc y)), (forall (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc z)), l_e_st_eq_landau_n_is y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc y)) => (fun (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc z)) => l_e_st_eq_landau_n_ax4 y z (l_e_tris1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_suc z) x i j)))))). Time Defined. (* constant 557 *) Definition l_e_st_eq_landau_n_23_t6 : (forall (x:l_e_st_eq_landau_n_nat), l_e_amone l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc y)) => (fun (v:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc z)) => l_e_st_eq_landau_n_23_t5 x y z u v))))). Time Defined. (* constant 558 *) Definition l_e_st_eq_landau_n_satz3a : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_one (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_onei l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_23_t6 x) (l_e_st_eq_landau_n_satz3 x n))). Time Defined. (* constant 559 *) Definition l_e_st_eq_landau_n_24_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_all (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (f y))))). Time Defined. (* constant 560 *) Definition l_e_st_eq_landau_n_24_prop2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_and (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x f))). Time Defined. (* constant 561 *) Definition l_e_st_eq_landau_n_24_prop3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), Prop)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (a y) (b y))))))). Time Defined. (* constant 562 *) Definition l_e_st_eq_landau_n_24_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), l_e_st_eq_landau_n_is (a l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => l_ande1 (l_e_st_eq_landau_n_is (a l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x a) pa))))). Time Defined. (* constant 563 *) Definition l_e_st_eq_landau_n_24_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), l_e_st_eq_landau_n_is (b l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => l_ande1 (l_e_st_eq_landau_n_is (b l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x b) pb))))). Time Defined. (* constant 564 *) Definition l_e_st_eq_landau_n_24_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), l_e_st_eq_landau_n_24_prop3 x a b pa pb l_e_st_eq_landau_n_1))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => l_e_tris2 l_e_st_eq_landau_n_nat (a l_e_st_eq_landau_n_1) (b l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_24_t1 x a b pa pb) (l_e_st_eq_landau_n_24_t2 x a b pa pb)))))). Time Defined. (* constant 565 *) Definition l_e_st_eq_landau_n_24_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (a y)) (l_e_st_eq_landau_n_suc (b y))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y) => l_e_st_eq_landau_n_ax2 (a y) (b y) p))))))). Time Defined. (* constant 566 *) Definition l_e_st_eq_landau_n_24_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y), l_e_st_eq_landau_n_24_prop1 x a))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y) => l_ande2 (l_e_st_eq_landau_n_is (a l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x a) pa))))))). Time Defined. (* constant 567 *) Definition l_e_st_eq_landau_n_24_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y), l_e_st_eq_landau_n_24_prop1 x b))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y) => l_ande2 (l_e_st_eq_landau_n_is (b l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x b) pb))))))). Time Defined. (* constant 568 *) Definition l_e_st_eq_landau_n_24_t7 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y), l_e_st_eq_landau_n_is (a (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (a y))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y) => l_e_st_eq_landau_n_24_t5 x a b pa pb y p y))))))). Time Defined. (* constant 569 *) Definition l_e_st_eq_landau_n_24_t8 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y), l_e_st_eq_landau_n_is (b (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (b y))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y) => l_e_st_eq_landau_n_24_t6 x a b pa pb y p y))))))). Time Defined. (* constant 570 *) Definition l_e_st_eq_landau_n_24_t9 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y), l_e_st_eq_landau_n_24_prop3 x a b pa pb (l_e_st_eq_landau_n_suc y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop3 x a b pa pb y) => l_e_tr3is l_e_st_eq_landau_n_nat (a (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (a y)) (l_e_st_eq_landau_n_suc (b y)) (b (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_24_t7 x a b pa pb y p) (l_e_st_eq_landau_n_24_t4 x a b pa pb y p) (l_e_symis l_e_st_eq_landau_n_nat (b (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (b y)) (l_e_st_eq_landau_n_24_t8 x a b pa pb y p))))))))). Time Defined. (* constant 571 *) Definition l_e_st_eq_landau_n_24_t10 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_24_prop3 x a b pa pb y)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_24_prop3 x a b pa pb z) (l_e_st_eq_landau_n_24_t3 x a b pa pb) (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_24_prop3 x a b pa pb z) => l_e_st_eq_landau_n_24_t9 x a b pa pb z u)) y)))))). Time Defined. (* constant 572 *) Definition l_e_st_eq_landau_n_24_t11 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_24_prop2 x a), (forall (pb:l_e_st_eq_landau_n_24_prop2 x b), l_e_is (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) a b))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_24_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_24_prop2 x b) => l_e_fisi l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat a b (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_24_t10 x a b pa pb y)))))). Time Defined. (* constant 573 *) Definition l_e_st_eq_landau_n_24_aa : (forall (x:l_e_st_eq_landau_n_nat), l_e_amone (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_24_prop2 x z)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (u:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (v:l_e_st_eq_landau_n_24_prop2 x z) => (fun (w:l_e_st_eq_landau_n_24_prop2 x u) => l_e_st_eq_landau_n_24_t11 x z u v w))))). Time Defined. (* constant 574 *) Definition l_e_st_eq_landau_n_24_prop4 : (forall (x:l_e_st_eq_landau_n_nat), Prop). exact (fun (x:l_e_st_eq_landau_n_nat) => l_some (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_24_prop2 x z)). Time Defined. (* constant 575 *) Definition l_e_st_eq_landau_n_24_t12 : l_e_st_eq_landau_n_24_prop1 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_suc. exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_refis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_suc x))). Time Defined. (* constant 576 *) Definition l_e_st_eq_landau_n_24_t13 : l_e_st_eq_landau_n_24_prop2 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_suc. exact (l_andi (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_24_prop1 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_suc) (l_e_refis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_24_t12). Time Defined. (* constant 577 *) Definition l_e_st_eq_landau_n_24_t14 : l_e_st_eq_landau_n_24_prop4 l_e_st_eq_landau_n_1. exact (l_somei (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_24_prop2 l_e_st_eq_landau_n_1 z) l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_24_t13). Time Defined. (* constant 578 *) Definition l_e_st_eq_landau_n_24_g : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_suc (f y)))))). Time Defined. (* constant 579 *) Definition l_e_st_eq_landau_n_24_t15 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_24_g x p f pf y) (l_e_st_eq_landau_n_suc (f y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_refis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_24_g x p f pf y)))))). Time Defined. (* constant 580 *) Definition l_e_st_eq_landau_n_24_t16 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => l_ande1 (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x f) pf)))). Time Defined. (* constant 581 *) Definition l_e_st_eq_landau_n_24_t17 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_24_g x p f pf l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_suc x)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_24_g x p f pf l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (f l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_t15 x p f pf l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ax2 (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_24_t16 x p f pf)))))). Time Defined. (* constant 582 *) Definition l_e_st_eq_landau_n_24_t18 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_24_prop1 x f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_ande2 (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x f) pf))))). Time Defined. (* constant 583 *) Definition l_e_st_eq_landau_n_24_t19 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (f y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_24_t18 x p f pf y y))))). Time Defined. (* constant 584 *) Definition l_e_st_eq_landau_n_24_t20 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_24_g x p f pf y)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris2 l_e_st_eq_landau_n_nat (f (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_24_g x p f pf y) (l_e_st_eq_landau_n_suc (f y)) (l_e_st_eq_landau_n_24_t19 x p f pf y) (l_e_st_eq_landau_n_24_t15 x p f pf y)))))). Time Defined. (* constant 585 *) Definition l_e_st_eq_landau_n_24_t21 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_24_g x p f pf (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_24_g x p f pf y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_24_g x p f pf (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (f (l_e_st_eq_landau_n_suc y))) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_24_g x p f pf y)) (l_e_st_eq_landau_n_24_t15 x p f pf (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_ax2 (f (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_24_g x p f pf y) (l_e_st_eq_landau_n_24_t20 x p f pf y))))))). Time Defined. (* constant 586 *) Definition l_e_st_eq_landau_n_24_t22 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), l_e_st_eq_landau_n_24_prop1 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_24_g x p f pf))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_24_t21 x p f pf y))))). Time Defined. (* constant 587 *) Definition l_e_st_eq_landau_n_24_t23 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), l_e_st_eq_landau_n_24_prop2 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_24_g x p f pf))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => l_andi (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_24_g x p f pf l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_suc x))) (l_e_st_eq_landau_n_24_prop1 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_24_g x p f pf)) (l_e_st_eq_landau_n_24_t17 x p f pf) (l_e_st_eq_landau_n_24_t22 x p f pf))))). Time Defined. (* constant 588 *) Definition l_e_st_eq_landau_n_24_t24 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_24_prop2 x f), l_e_st_eq_landau_n_24_prop4 (l_e_st_eq_landau_n_suc x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_24_prop2 x f) => l_somei (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_24_prop2 (l_e_st_eq_landau_n_suc x) z) (l_e_st_eq_landau_n_24_g x p f pf) (l_e_st_eq_landau_n_24_t23 x p f pf))))). Time Defined. (* constant 589 *) Definition l_e_st_eq_landau_n_24_t25 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_24_prop4 x), l_e_st_eq_landau_n_24_prop4 (l_e_st_eq_landau_n_suc x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_24_prop4 x) => l_someapp (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_24_prop2 x z) p (l_e_st_eq_landau_n_24_prop4 (l_e_st_eq_landau_n_suc x)) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (u:l_e_st_eq_landau_n_24_prop2 x z) => l_e_st_eq_landau_n_24_t24 x p z u)))). Time Defined. (* constant 590 *) Definition l_e_st_eq_landau_n_24_bb : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_24_prop4 x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_24_prop4 y) l_e_st_eq_landau_n_24_t14 (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_24_prop4 y) => l_e_st_eq_landau_n_24_t25 y u)) x). Time Defined. (* constant 591 *) Definition l_e_st_eq_landau_n_satz4 : (forall (x:l_e_st_eq_landau_n_nat), l_e_one (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_and (l_e_st_eq_landau_n_is (z l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_all (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (z (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (z y)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_onei (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_24_prop2 x z) (l_e_st_eq_landau_n_24_aa x) (l_e_st_eq_landau_n_24_bb x)). Time Defined. (* constant 592 *) Definition l_e_st_eq_landau_n_plus : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_ind (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_24_prop2 x z) (l_e_st_eq_landau_n_satz4 x)). Time Defined. (* constant 593 *) Definition l_e_st_eq_landau_n_pl : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_plus x y)). Time Defined. (* constant 594 *) Definition l_e_st_eq_landau_n_24_t26 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_24_prop2 x (l_e_st_eq_landau_n_plus x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_oneax (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_24_prop2 x z) (l_e_st_eq_landau_n_satz4 x)). Time Defined. (* constant 595 *) Definition l_e_st_eq_landau_n_satz4a : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_ande1 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_plus x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x (l_e_st_eq_landau_n_plus x)) (l_e_st_eq_landau_n_24_t26 x)). Time Defined. (* constant 596 *) Definition l_e_st_eq_landau_n_24_t27 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_24_prop1 x (l_e_st_eq_landau_n_plus x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_ande2 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_plus x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_prop1 x (l_e_st_eq_landau_n_plus x)) (l_e_st_eq_landau_n_24_t26 x)). Time Defined. (* constant 597 *) Definition l_e_st_eq_landau_n_satz4b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_24_t27 x y)). Time Defined. (* constant 598 *) Definition l_e_st_eq_landau_n_24_t28 : l_e_is (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (l_e_st_eq_landau_n_plus l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_suc. exact (l_e_st_eq_landau_n_24_t11 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_plus l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_24_t26 l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_24_t13). Time Defined. (* constant 599 *) Definition l_e_st_eq_landau_n_satz4c : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 x) (l_e_st_eq_landau_n_suc x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_fise l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_plus l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_24_t28 x). Time Defined. (* constant 600 *) Definition l_e_st_eq_landau_n_24_t29 : (forall (x:l_e_st_eq_landau_n_nat), l_e_is (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (l_e_st_eq_landau_n_plus (l_e_st_eq_landau_n_suc x)) (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_plus x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_24_t11 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_plus (l_e_st_eq_landau_n_suc x)) (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_plus x y)) (l_e_st_eq_landau_n_24_t26 (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_24_t23 x (l_e_st_eq_landau_n_24_bb x) (l_e_st_eq_landau_n_plus x) (l_e_st_eq_landau_n_24_t26 x))). Time Defined. (* constant 601 *) Definition l_e_st_eq_landau_n_satz4d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_fise l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_plus (l_e_st_eq_landau_n_suc x)) (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_plus x z)) (l_e_st_eq_landau_n_24_t29 x) y)). Time Defined. (* constant 602 *) Definition l_e_st_eq_landau_n_satz4e : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_satz4a x)). Time Defined. (* constant 603 *) Definition l_e_st_eq_landau_n_satz4f : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_satz4b x y))). Time Defined. (* constant 604 *) Definition l_e_st_eq_landau_n_satz4g : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 x) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_satz4c x)). Time Defined. (* constant 605 *) Definition l_e_st_eq_landau_n_satz4h : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc x) y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_satz4d x y))). Time Defined. (* constant 606 *) Definition l_e_st_eq_landau_n_ispl1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl u z) x y i)))). Time Defined. (* constant 607 *) Definition l_e_st_eq_landau_n_ispl2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl z u) x y i)))). Time Defined. (* constant 608 *) Definition l_e_st_eq_landau_n_ispl12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_ispl1 x y z i) (l_e_st_eq_landau_n_ispl2 z u y j))))))). Time Defined. (* constant 609 *) Definition l_e_st_eq_landau_n_25_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), Prop))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y z))))). Time Defined. (* constant 610 *) Definition l_e_st_eq_landau_n_25_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_25_prop1 x y l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_satz4a (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_satz4f x y) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_satz4e y)))). Time Defined. (* constant 611 *) Definition l_e_st_eq_landau_n_25_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_25_prop1 x y z), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) z)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_25_prop1 x y z) => l_e_st_eq_landau_n_ax2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y z)) p)))). Time Defined. (* constant 612 *) Definition l_e_st_eq_landau_n_25_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_25_prop1 x y z), l_e_st_eq_landau_n_25_prop1 x y (l_e_st_eq_landau_n_suc z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_25_prop1 x y z) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_suc z)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) z)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y z))) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl y z))) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc z))) (l_e_st_eq_landau_n_satz4b (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_25_t2 x y z p) (l_e_st_eq_landau_n_satz4f x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc z)) x (l_e_st_eq_landau_n_satz4f y z)))))). Time Defined. (* constant 613 *) Definition l_e_st_eq_landau_n_satz5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_25_prop1 x y u) (l_e_st_eq_landau_n_25_t1 x y) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_25_prop1 x y u) => l_e_st_eq_landau_n_25_t3 x y u v)) z))). Time Defined. (* constant 614 *) Definition l_e_st_eq_landau_n_asspl1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz5 x y z))). Time Defined. (* constant 615 *) Definition l_e_st_eq_landau_n_asspl2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) z)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_satz5 x y z)))). Time Defined. (* constant 616 *) Definition l_e_st_eq_landau_n_26_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl y x))). Time Defined. (* constant 617 *) Definition l_e_st_eq_landau_n_26_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz4a y)). Time Defined. (* constant 618 *) Definition l_e_st_eq_landau_n_26_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 y) (l_e_st_eq_landau_n_suc y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz4c y)). Time Defined. (* constant 619 *) Definition l_e_st_eq_landau_n_26_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_26_prop1 l_e_st_eq_landau_n_1 y)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris2 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 y) (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_26_t2 x y) (l_e_st_eq_landau_n_26_t1 x y))). Time Defined. (* constant 620 *) Definition l_e_st_eq_landau_n_26_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_26_prop1 x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_26_prop1 x y) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl y x)) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_ax2 (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl y x) p) (l_e_st_eq_landau_n_satz4f y x)))). Time Defined. (* constant 621 *) Definition l_e_st_eq_landau_n_26_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_26_prop1 x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_26_prop1 x y) => l_e_st_eq_landau_n_satz4d x y))). Time Defined. (* constant 622 *) Definition l_e_st_eq_landau_n_26_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_26_prop1 x y), l_e_st_eq_landau_n_26_prop1 (l_e_st_eq_landau_n_suc x) y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_26_prop1 x y) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_26_t5 x y p) (l_e_st_eq_landau_n_26_t4 x y p)))). Time Defined. (* constant 623 *) Definition l_e_st_eq_landau_n_satz6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_26_prop1 z y) (l_e_st_eq_landau_n_26_t3 x y) (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_26_prop1 z y) => l_e_st_eq_landau_n_26_t6 z y u)) x)). Time Defined. (* constant 624 *) Definition l_e_st_eq_landau_n_compl : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz6 x y)). Time Defined. (* constant 625 *) Definition l_e_st_eq_landau_n_26_t7 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_26_prop1 x l_e_st_eq_landau_n_1). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 x) (l_e_st_eq_landau_n_satz4a x) (l_e_st_eq_landau_n_satz4g x)). Time Defined. (* constant 626 *) Definition l_e_st_eq_landau_n_26_t8 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_26_prop1 x y), l_e_st_eq_landau_n_26_prop1 x (l_e_st_eq_landau_n_suc y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_26_prop1 x y) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl y x)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc y) x) (l_e_st_eq_landau_n_satz4b x y) (l_e_st_eq_landau_n_ax2 (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl y x) p) (l_e_st_eq_landau_n_satz4h y x)))). Time Defined. (* constant 627 *) Definition l_e_st_eq_landau_n_26_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_26_prop1 x z) (l_e_st_eq_landau_n_26_t7 x) (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_26_prop1 x z) => l_e_st_eq_landau_n_26_t8 x z u)) y)). Time Defined. (* constant 628 *) Definition l_e_st_eq_landau_n_27_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_nis y (l_e_st_eq_landau_n_pl x y))). Time Defined. (* constant 629 *) Definition l_e_st_eq_landau_n_27_t1 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nis l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_suc x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_symnotis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_ax3 x)). Time Defined. (* constant 630 *) Definition l_e_st_eq_landau_n_27_t2 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_27_prop1 x l_e_st_eq_landau_n_1). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_notis_th4 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_27_t1 x) (l_e_st_eq_landau_n_satz4a x)). Time Defined. (* constant 631 *) Definition l_e_st_eq_landau_n_27_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_27_prop1 x y), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_27_prop1 x y) => l_e_st_eq_landau_n_satz1 y (l_e_st_eq_landau_n_pl x y) p))). Time Defined. (* constant 632 *) Definition l_e_st_eq_landau_n_27_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_27_prop1 x y), l_e_st_eq_landau_n_27_prop1 x (l_e_st_eq_landau_n_suc y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_27_prop1 x y) => l_e_notis_th4 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_27_t3 x y p) (l_e_st_eq_landau_n_satz4b x y)))). Time Defined. (* constant 633 *) Definition l_e_st_eq_landau_n_satz7 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nis y (l_e_st_eq_landau_n_pl x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_27_prop1 x z) (l_e_st_eq_landau_n_27_t2 x) (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_27_prop1 x z) => l_e_st_eq_landau_n_27_t4 x z u)) y)). Time Defined. (* constant 634 *) Definition l_e_st_eq_landau_n_28_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis y z), Prop)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis y z) => l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x z))))). Time Defined. (* constant 635 *) Definition l_e_st_eq_landau_n_28_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis y z), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_suc z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis y z) => l_e_st_eq_landau_n_satz1 y z n)))). Time Defined. (* constant 636 *) Definition l_e_st_eq_landau_n_28_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis y z), l_e_st_eq_landau_n_28_prop1 l_e_st_eq_landau_n_1 y z n)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis y z) => l_e_notis_th5 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_suc z) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 y) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 z) (l_e_st_eq_landau_n_28_t1 x y z n) (l_e_st_eq_landau_n_satz4g y) (l_e_st_eq_landau_n_satz4g z))))). Time Defined. (* constant 637 *) Definition l_e_st_eq_landau_n_28_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis y z), (forall (p:l_e_st_eq_landau_n_28_prop1 x y z n), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis y z) => (fun (p:l_e_st_eq_landau_n_28_prop1 x y z n) => l_e_st_eq_landau_n_satz1 (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x z) p))))). Time Defined. (* constant 638 *) Definition l_e_st_eq_landau_n_28_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis y z), (forall (p:l_e_st_eq_landau_n_28_prop1 x y z n), l_e_st_eq_landau_n_28_prop1 (l_e_st_eq_landau_n_suc x) y z n))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis y z) => (fun (p:l_e_st_eq_landau_n_28_prop1 x y z n) => l_e_notis_th5 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc x) z) (l_e_st_eq_landau_n_28_t3 x y z n p) (l_e_st_eq_landau_n_satz4h x y) (l_e_st_eq_landau_n_satz4h x z)))))). Time Defined. (* constant 639 *) Definition l_e_st_eq_landau_n_satz8 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis y z), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis y z) => l_e_st_eq_landau_n_induction (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_28_prop1 u y z n) (l_e_st_eq_landau_n_28_t2 x y z n) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_28_prop1 u y z n) => l_e_st_eq_landau_n_28_t4 u y z n v)) x)))). Time Defined. (* constant 640 *) Definition l_e_st_eq_landau_n_satz8a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x z)), l_e_st_eq_landau_n_is y z)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x z)) => l_imp_th7 (l_e_st_eq_landau_n_is y z) (l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x z)) (l_weli (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x z)) i) (fun (u:l_e_st_eq_landau_n_nis y z) => l_e_st_eq_landau_n_satz8 x y z u))))). Time Defined. (* constant 641 *) Definition l_e_st_eq_landau_n_diffprop : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), Prop))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl y z)))). Time Defined. (* constant 642 *) Definition l_e_st_eq_landau_n_28_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (v:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (dv:l_e_st_eq_landau_n_diffprop x y v), l_e_st_eq_landau_n_is u v)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (dv:l_e_st_eq_landau_n_diffprop x y v) => l_e_st_eq_landau_n_satz8a y u v (l_e_tris1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl y v) x du dv))))))). Time Defined. (* constant 643 *) Definition l_e_st_eq_landau_n_satz8b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_amone l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl y z)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (dv:l_e_st_eq_landau_n_diffprop x y v) => l_e_st_eq_landau_n_28_t5 x y u v du dv)))))). Time Defined. (* constant 644 *) Definition l_e_st_eq_landau_n_29_i : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x y)). Time Defined. (* constant 645 *) Definition l_e_st_eq_landau_n_29_ii : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u))). Time Defined. (* constant 646 *) Definition l_e_st_eq_landau_n_29_iii : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_some (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop y x v))). Time Defined. (* constant 647 *) Definition l_e_st_eq_landau_n_29_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (one1:l_e_st_eq_landau_n_29_i x y), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl u x) (l_e_st_eq_landau_n_pl y u))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (one1:l_e_st_eq_landau_n_29_i x y) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl u x) (l_e_st_eq_landau_n_pl x u) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_compl u x) (l_e_st_eq_landau_n_ispl1 x y u one1))))). Time Defined. (* constant 648 *) Definition l_e_st_eq_landau_n_29_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (one1:l_e_st_eq_landau_n_29_i x y), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nis x (l_e_st_eq_landau_n_pl y u))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (one1:l_e_st_eq_landau_n_29_i x y) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_notis_th3 l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl u x) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_satz7 u x) (l_e_st_eq_landau_n_29_t1 x y one1 u))))). Time Defined. (* constant 649 *) Definition l_e_st_eq_landau_n_29_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (one1:l_e_st_eq_landau_n_29_i x y), l_not (l_e_st_eq_landau_n_29_ii x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (one1:l_e_st_eq_landau_n_29_i x y) => l_some_th5 l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u) (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_29_t2 x y one1 u)))). Time Defined. (* constant 650 *) Definition l_e_st_eq_landau_n_29_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_ec (l_e_st_eq_landau_n_29_i x y) (l_e_st_eq_landau_n_29_ii x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_ec_th1 (l_e_st_eq_landau_n_29_i x y) (l_e_st_eq_landau_n_29_ii x y) (fun (z:l_e_st_eq_landau_n_29_i x y) => l_e_st_eq_landau_n_29_t3 x y z))). Time Defined. (* constant 651 *) Definition l_e_st_eq_landau_n_29_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (one1:l_e_st_eq_landau_n_29_i x y), l_not (l_e_st_eq_landau_n_29_iii x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (one1:l_e_st_eq_landau_n_29_i x y) => l_e_st_eq_landau_n_29_t3 y x (l_e_symis l_e_st_eq_landau_n_nat x y one1)))). Time Defined. (* constant 652 *) Definition l_e_st_eq_landau_n_29_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_ec (l_e_st_eq_landau_n_29_iii x y) (l_e_st_eq_landau_n_29_i x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_ec_th2 (l_e_st_eq_landau_n_29_iii x y) (l_e_st_eq_landau_n_29_i x y) (fun (z:l_e_st_eq_landau_n_29_i x y) => l_e_st_eq_landau_n_29_t5 x y z))). Time Defined. (* constant 653 *) Definition l_e_st_eq_landau_n_29_t6a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (three1:l_e_st_eq_landau_n_29_iii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (v:l_e_st_eq_landau_n_nat), (forall (dv:l_e_st_eq_landau_n_diffprop y x v), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl v u) x))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (three1:l_e_st_eq_landau_n_29_iii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => l_e_tr4is l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x v) u) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl v u)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl v u) x) du (l_e_st_eq_landau_n_ispl1 y (l_e_st_eq_landau_n_pl x v) u dv) (l_e_st_eq_landau_n_asspl1 x v u) (l_e_st_eq_landau_n_compl x (l_e_st_eq_landau_n_pl v u)))))))))). Time Defined. (* constant 654 *) Definition l_e_st_eq_landau_n_29_t7 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (three1:l_e_st_eq_landau_n_29_iii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (v:l_e_st_eq_landau_n_nat), (forall (dv:l_e_st_eq_landau_n_diffprop y x v), l_con)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (three1:l_e_st_eq_landau_n_29_iii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => l_mp (l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl v u) x)) l_con (l_e_st_eq_landau_n_29_t6a x y two1 three1 u du v dv) (l_e_st_eq_landau_n_satz7 (l_e_st_eq_landau_n_pl v u) x))))))))). Time Defined. (* constant 655 *) Definition l_e_st_eq_landau_n_29_t8 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (three1:l_e_st_eq_landau_n_29_iii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), l_con)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (three1:l_e_st_eq_landau_n_29_iii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_someapp l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop y x v) three1 l_con (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => l_e_st_eq_landau_n_29_t7 x y two1 three1 u du v dv)))))))). Time Defined. (* constant 656 *) Definition l_e_st_eq_landau_n_29_t9 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (three1:l_e_st_eq_landau_n_29_iii x y), l_con)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (three1:l_e_st_eq_landau_n_29_iii x y) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u) two1 l_con (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_e_st_eq_landau_n_29_t8 x y two1 three1 u du)))))). Time Defined. (* constant 657 *) Definition l_e_st_eq_landau_n_29_t10 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (two1:l_e_st_eq_landau_n_29_ii x y), l_not (l_e_st_eq_landau_n_29_iii x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (z:l_e_st_eq_landau_n_29_iii x y) => l_e_st_eq_landau_n_29_t9 x y two1 z)))). Time Defined. (* constant 658 *) Definition l_e_st_eq_landau_n_29_t11 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_ec (l_e_st_eq_landau_n_29_ii x y) (l_e_st_eq_landau_n_29_iii x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_ec_th1 (l_e_st_eq_landau_n_29_ii x y) (l_e_st_eq_landau_n_29_iii x y) (fun (z:l_e_st_eq_landau_n_29_ii x y) => l_e_st_eq_landau_n_29_t10 x y z))). Time Defined. (* constant 659 *) Definition l_e_st_eq_landau_n_29_a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_ec3 (l_e_st_eq_landau_n_29_i x y) (l_e_st_eq_landau_n_29_ii x y) (l_e_st_eq_landau_n_29_iii x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_ec3_th6 (l_e_st_eq_landau_n_29_i x y) (l_e_st_eq_landau_n_29_ii x y) (l_e_st_eq_landau_n_29_iii x y) (l_e_st_eq_landau_n_29_t4 x y) (l_e_st_eq_landau_n_29_t11 x y) (l_e_st_eq_landau_n_29_t6 x y))). Time Defined. (* constant 660 *) Definition l_e_st_eq_landau_n_29_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_or3 (l_e_st_eq_landau_n_29_i x y) (l_e_st_eq_landau_n_29_ii x y) (l_e_st_eq_landau_n_29_iii x y))). Time Defined. (* constant 661 *) Definition l_e_st_eq_landau_n_29_t12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_29_prop1 x l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_or3i1 (l_e_st_eq_landau_n_29_i x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_29_ii x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_29_iii x l_e_st_eq_landau_n_1) j)). Time Defined. (* constant 662 *) Definition l_e_st_eq_landau_n_29_t13 : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 u))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) => l_e_tris l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_suc u) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 u) j (l_e_st_eq_landau_n_satz4g u))))). Time Defined. (* constant 663 *) Definition l_e_st_eq_landau_n_29_t14 : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)), l_e_st_eq_landau_n_29_ii x l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) => l_somei l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x l_e_st_eq_landau_n_1 z) u (l_e_st_eq_landau_n_29_t13 x n u j))))). Time Defined. (* constant 664 *) Definition l_e_st_eq_landau_n_29_t15 : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)), l_e_st_eq_landau_n_29_ii x l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) => l_someapp l_e_st_eq_landau_n_nat (fun (u1:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u1)) (l_e_st_eq_landau_n_satz3 x n) (l_e_st_eq_landau_n_29_ii x l_e_st_eq_landau_n_1) (fun (u1:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u1)) => l_e_st_eq_landau_n_29_t14 x n u1 z)))))). Time Defined. (* constant 665 *) Definition l_e_st_eq_landau_n_29_t16 : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)), l_e_st_eq_landau_n_29_prop1 x l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) => l_or3i2 (l_e_st_eq_landau_n_29_i x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_29_ii x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_29_iii x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_29_t15 x n u j))))). Time Defined. (* constant 666 *) Definition l_e_st_eq_landau_n_29_t16a : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_29_prop1 x l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_satz3 x n) (l_e_st_eq_landau_n_29_prop1 x l_e_st_eq_landau_n_1) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) => l_e_st_eq_landau_n_29_t16 x n u v)))). Time Defined. (* constant 667 *) Definition l_e_st_eq_landau_n_29_t17 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_29_prop1 x l_e_st_eq_landau_n_1). exact (fun (x:l_e_st_eq_landau_n_nat) => l_imp_th1 (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_29_prop1 x l_e_st_eq_landau_n_1) (fun (z:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_29_t12 x z) (fun (z:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_29_t16a x z)). Time Defined. (* constant 668 *) Definition l_e_st_eq_landau_n_29_t18 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (one1:l_e_st_eq_landau_n_29_i x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (one1:l_e_st_eq_landau_n_29_i x y) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz4e y) (l_e_st_eq_landau_n_ispl1 y x l_e_st_eq_landau_n_1 (l_e_symis l_e_st_eq_landau_n_nat x y one1)))))). Time Defined. (* constant 669 *) Definition l_e_st_eq_landau_n_29_t19 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (one1:l_e_st_eq_landau_n_29_i x y), l_e_st_eq_landau_n_29_iii x (l_e_st_eq_landau_n_suc y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (one1:l_e_st_eq_landau_n_29_i x y) => l_somei l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_suc y) x z) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_29_t18 x y p one1))))). Time Defined. (* constant 670 *) Definition l_e_st_eq_landau_n_29_t20 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (one1:l_e_st_eq_landau_n_29_i x y), l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (one1:l_e_st_eq_landau_n_29_i x y) => l_or3i3 (l_e_st_eq_landau_n_29_i x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_ii x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_iii x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_t19 x y p one1))))). Time Defined. (* constant 671 *) Definition l_e_st_eq_landau_n_29_t21 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (j:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (j:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_e_tr3is l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc y) du (l_e_st_eq_landau_n_ispl2 u l_e_st_eq_landau_n_1 y j) (l_e_st_eq_landau_n_satz4a y)))))))). Time Defined. (* constant 672 *) Definition l_e_st_eq_landau_n_29_t22 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (j:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (j:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_or3i1 (l_e_st_eq_landau_n_29_i x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_ii x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_iii x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_t21 x y p two1 u du j)))))))). Time Defined. (* constant 673 *) Definition l_e_st_eq_landau_n_29_t23 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1), (forall (w:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_suc w)), l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 w)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1) => (fun (w:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_suc w)) => l_e_tris l_e_st_eq_landau_n_nat u (l_e_st_eq_landau_n_suc w) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 w) j (l_e_st_eq_landau_n_satz4g w)))))))))). Time Defined. (* constant 674 *) Definition l_e_st_eq_landau_n_29_t24 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1), (forall (w:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_suc w)), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc y) w)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1) => (fun (w:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_suc w)) => l_e_tr4is l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) w) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc y) w) du (l_e_st_eq_landau_n_ispl2 u (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 w) y (l_e_st_eq_landau_n_29_t23 x y p two1 u du n w j)) (l_e_st_eq_landau_n_asspl2 y l_e_st_eq_landau_n_1 w) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc y) w (l_e_st_eq_landau_n_satz4a y))))))))))). Time Defined. (* constant 675 *) Definition l_e_st_eq_landau_n_29_t25 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1), (forall (w:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_suc w)), l_e_st_eq_landau_n_29_ii x (l_e_st_eq_landau_n_suc y)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1) => (fun (w:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_suc w)) => l_somei l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x (l_e_st_eq_landau_n_suc y) z) w (l_e_st_eq_landau_n_29_t24 x y p two1 u du n w j)))))))))). Time Defined. (* constant 676 *) Definition l_e_st_eq_landau_n_29_t26 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_29_ii x (l_e_st_eq_landau_n_suc y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1) => l_someapp l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_suc z)) (l_e_st_eq_landau_n_satz3 u n) (l_e_st_eq_landau_n_29_ii x (l_e_st_eq_landau_n_suc y)) (fun (z:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_is u (l_e_st_eq_landau_n_suc z)) => l_e_st_eq_landau_n_29_t25 x y p two1 u du n z v))))))))). Time Defined. (* constant 677 *) Definition l_e_st_eq_landau_n_29_t27 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), (forall (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => (fun (n:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1) => l_or3i2 (l_e_st_eq_landau_n_29_i x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_ii x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_iii x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_t26 x y p two1 u du n)))))))). Time Defined. (* constant 678 *) Definition l_e_st_eq_landau_n_29_t28 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_imp_th1 (l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y)) (fun (z:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_29_t22 x y p two1 u du z) (fun (z:l_e_st_eq_landau_n_nis u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_29_t27 x y p two1 u du z))))))). Time Defined. (* constant 679 *) Definition l_e_st_eq_landau_n_29_t28a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (two1:l_e_st_eq_landau_n_29_ii x y), l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (two1:l_e_st_eq_landau_n_29_ii x y) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u) two1 (l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y)) (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_e_st_eq_landau_n_29_t28 x y p two1 u du)))))). Time Defined. (* constant 680 *) Definition l_e_st_eq_landau_n_29_t29 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (three1:l_e_st_eq_landau_n_29_iii x y), (forall (v:l_e_st_eq_landau_n_nat), (forall (dv:l_e_st_eq_landau_n_diffprop y x v), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc v)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (three1:l_e_st_eq_landau_n_29_iii x y) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x v)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc v)) (l_e_st_eq_landau_n_ax2 y (l_e_st_eq_landau_n_pl x v) dv) (l_e_st_eq_landau_n_satz4f x v))))))). Time Defined. (* constant 681 *) Definition l_e_st_eq_landau_n_29_t30 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (three1:l_e_st_eq_landau_n_29_iii x y), (forall (v:l_e_st_eq_landau_n_nat), (forall (dv:l_e_st_eq_landau_n_diffprop y x v), l_e_st_eq_landau_n_29_iii x (l_e_st_eq_landau_n_suc y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (three1:l_e_st_eq_landau_n_29_iii x y) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => l_somei l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_suc y) x z) (l_e_st_eq_landau_n_suc v) (l_e_st_eq_landau_n_29_t29 x y p three1 v dv))))))). Time Defined. (* constant 682 *) Definition l_e_st_eq_landau_n_29_t31 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (three1:l_e_st_eq_landau_n_29_iii x y), l_e_st_eq_landau_n_29_iii x (l_e_st_eq_landau_n_suc y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (three1:l_e_st_eq_landau_n_29_iii x y) => l_someapp l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop y x v) three1 (l_e_st_eq_landau_n_29_iii x (l_e_st_eq_landau_n_suc y)) (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => l_e_st_eq_landau_n_29_t30 x y p three1 v dv)))))). Time Defined. (* constant 683 *) Definition l_e_st_eq_landau_n_29_t32 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), (forall (three1:l_e_st_eq_landau_n_29_iii x y), l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => (fun (three1:l_e_st_eq_landau_n_29_iii x y) => l_or3i3 (l_e_st_eq_landau_n_29_i x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_ii x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_iii x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_29_t31 x y p three1))))). Time Defined. (* constant 684 *) Definition l_e_st_eq_landau_n_29_t33 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_29_prop1 x y), l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_29_prop1 x y) => l_or3app (l_e_st_eq_landau_n_29_i x y) (l_e_st_eq_landau_n_29_ii x y) (l_e_st_eq_landau_n_29_iii x y) (l_e_st_eq_landau_n_29_prop1 x (l_e_st_eq_landau_n_suc y)) p (fun (z:l_e_st_eq_landau_n_29_i x y) => l_e_st_eq_landau_n_29_t20 x y p z) (fun (z:l_e_st_eq_landau_n_29_ii x y) => l_e_st_eq_landau_n_29_t28a x y p z) (fun (z:l_e_st_eq_landau_n_29_iii x y) => l_e_st_eq_landau_n_29_t32 x y p z)))). Time Defined. (* constant 685 *) Definition l_e_st_eq_landau_n_29_b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_or3 (l_e_st_eq_landau_n_29_i x y) (l_e_st_eq_landau_n_29_ii x y) (l_e_st_eq_landau_n_29_iii x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_29_prop1 x z) (l_e_st_eq_landau_n_29_t17 x) (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_29_prop1 x z) => l_e_st_eq_landau_n_29_t33 x z u)) y)). Time Defined. (* constant 686 *) Definition l_e_st_eq_landau_n_satz9 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_orec3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl y u))) (l_e_st_eq_landau_n_some (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is y (l_e_st_eq_landau_n_pl x v))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_orec3i (l_e_st_eq_landau_n_29_i x y) (l_e_st_eq_landau_n_29_ii x y) (l_e_st_eq_landau_n_29_iii x y) (l_e_st_eq_landau_n_29_b x y) (l_e_st_eq_landau_n_29_a x y))). Time Defined. (* constant 687 *) Definition l_e_st_eq_landau_n_satz9a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_or3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u)) (l_e_st_eq_landau_n_some (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop y x v)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_29_b x y)). Time Defined. (* constant 688 *) Definition l_e_st_eq_landau_n_satz9b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_ec3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u)) (l_e_st_eq_landau_n_some (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop y x v)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_29_a x y)). Time Defined. (* constant 689 *) Definition l_e_st_eq_landau_n_more : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u))). Time Defined. (* constant 690 *) Definition l_e_st_eq_landau_n_less : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_some (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop y x v))). Time Defined. (* constant 691 *) Definition l_e_st_eq_landau_n_satz10 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_orec3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz9 x y)). Time Defined. (* constant 692 *) Definition l_e_st_eq_landau_n_satz10a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_or3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz9a x y)). Time Defined. (* constant 693 *) Definition l_e_st_eq_landau_n_satz10b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_ec3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz9b x y)). Time Defined. (* constant 694 *) Definition l_e_st_eq_landau_n_satz11 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_less y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => m))). Time Defined. (* constant 695 *) Definition l_e_st_eq_landau_n_satz12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_more y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l))). Time Defined. (* constant 696 *) Definition l_e_st_eq_landau_n_moreis : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_or (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y))). Time Defined. (* constant 697 *) Definition l_e_st_eq_landau_n_lessis : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_or (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y))). Time Defined. (* constant 698 *) Definition l_e_st_eq_landau_n_satz13 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), l_e_st_eq_landau_n_lessis y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => l_or_th9 (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_less y x) (l_e_st_eq_landau_n_is y x) m (fun (z:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz11 x y z) (fun (z:l_e_st_eq_landau_n_is x y) => l_e_symis l_e_st_eq_landau_n_nat x y z)))). Time Defined. (* constant 699 *) Definition l_e_st_eq_landau_n_satz14 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), l_e_st_eq_landau_n_moreis y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => l_or_th9 (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more y x) (l_e_st_eq_landau_n_is y x) l (fun (z:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz12 x y z) (fun (z:l_e_st_eq_landau_n_is x y) => l_e_symis l_e_st_eq_landau_n_nat x y z)))). Time Defined. (* constant 700 *) Definition l_e_st_eq_landau_n_ismore1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (m:l_e_st_eq_landau_n_more x z), l_e_st_eq_landau_n_more y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (m:l_e_st_eq_landau_n_more x z) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_more u z) x y m i))))). Time Defined. (* constant 701 *) Definition l_e_st_eq_landau_n_ismore2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (m:l_e_st_eq_landau_n_more z x), l_e_st_eq_landau_n_more z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (m:l_e_st_eq_landau_n_more z x) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_more z u) x y m i))))). Time Defined. (* constant 702 *) Definition l_e_st_eq_landau_n_isless1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (l:l_e_st_eq_landau_n_less x z), l_e_st_eq_landau_n_less y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (l:l_e_st_eq_landau_n_less x z) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_less u z) x y l i))))). Time Defined. (* constant 703 *) Definition l_e_st_eq_landau_n_isless2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (l:l_e_st_eq_landau_n_less z x), l_e_st_eq_landau_n_less z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (l:l_e_st_eq_landau_n_less z x) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_less z u) x y l i))))). Time Defined. (* constant 704 *) Definition l_e_st_eq_landau_n_ismoreis1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (m:l_e_st_eq_landau_n_moreis x z), l_e_st_eq_landau_n_moreis y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (m:l_e_st_eq_landau_n_moreis x z) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_moreis u z) x y m i))))). Time Defined. (* constant 705 *) Definition l_e_st_eq_landau_n_ismoreis2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (m:l_e_st_eq_landau_n_moreis z x), l_e_st_eq_landau_n_moreis z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (m:l_e_st_eq_landau_n_moreis z x) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_moreis z u) x y m i))))). Time Defined. (* constant 706 *) Definition l_e_st_eq_landau_n_islessis1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (l:l_e_st_eq_landau_n_lessis x z), l_e_st_eq_landau_n_lessis y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (l:l_e_st_eq_landau_n_lessis x z) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis u z) x y l i))))). Time Defined. (* constant 707 *) Definition l_e_st_eq_landau_n_islessis2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (l:l_e_st_eq_landau_n_lessis z x), l_e_st_eq_landau_n_lessis z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (l:l_e_st_eq_landau_n_lessis z x) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis z u) x y l i))))). Time Defined. (* constant 708 *) Definition l_e_st_eq_landau_n_moreisi2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_moreis x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_ori2 (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) i))). Time Defined. (* constant 709 *) Definition l_e_st_eq_landau_n_lessisi2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_lessis x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_ori2 (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) i))). Time Defined. (* constant 710 *) Definition l_e_st_eq_landau_n_moreisi1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moreis x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_ori1 (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) m))). Time Defined. (* constant 711 *) Definition l_e_st_eq_landau_n_lessisi1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_lessis x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_ori1 (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) l))). Time Defined. (* constant 712 *) Definition l_e_st_eq_landau_n_ismore12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), (forall (m:l_e_st_eq_landau_n_more x z), l_e_st_eq_landau_n_more y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => (fun (m:l_e_st_eq_landau_n_more x z) => l_e_st_eq_landau_n_ismore2 z u y j (l_e_st_eq_landau_n_ismore1 x y z i m)))))))). Time Defined. (* constant 713 *) Definition l_e_st_eq_landau_n_isless12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), (forall (l:l_e_st_eq_landau_n_less x z), l_e_st_eq_landau_n_less y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => (fun (l:l_e_st_eq_landau_n_less x z) => l_e_st_eq_landau_n_isless2 z u y j (l_e_st_eq_landau_n_isless1 x y z i l)))))))). Time Defined. (* constant 714 *) Definition l_e_st_eq_landau_n_ismoreis12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), (forall (m:l_e_st_eq_landau_n_moreis x z), l_e_st_eq_landau_n_moreis y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => (fun (m:l_e_st_eq_landau_n_moreis x z) => l_e_st_eq_landau_n_ismoreis2 z u y j (l_e_st_eq_landau_n_ismoreis1 x y z i m)))))))). Time Defined. (* constant 715 *) Definition l_e_st_eq_landau_n_islessis12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), (forall (l:l_e_st_eq_landau_n_lessis x z), l_e_st_eq_landau_n_lessis y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => (fun (l:l_e_st_eq_landau_n_lessis x z) => l_e_st_eq_landau_n_islessis2 z u y j (l_e_st_eq_landau_n_islessis1 x y z i l)))))))). Time Defined. (* constant 716 *) Definition l_e_st_eq_landau_n_satz10c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), l_not (l_e_st_eq_landau_n_less x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => l_ec3_th7 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10b x y) (l_comor (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) m)))). Time Defined. (* constant 717 *) Definition l_e_st_eq_landau_n_satz10d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), l_not (l_e_st_eq_landau_n_more x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => l_ec3_th9 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10b x y) l))). Time Defined. (* constant 718 *) Definition l_e_st_eq_landau_n_satz10e : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (n:l_not (l_e_st_eq_landau_n_more x y)), l_e_st_eq_landau_n_lessis x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (n:l_not (l_e_st_eq_landau_n_more x y)) => l_or3_th2 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10a x y) n))). Time Defined. (* constant 719 *) Definition l_e_st_eq_landau_n_satz10f : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (n:l_not (l_e_st_eq_landau_n_less x y)), l_e_st_eq_landau_n_moreis x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (n:l_not (l_e_st_eq_landau_n_less x y)) => l_comor (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_or3_th3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10a x y) n)))). Time Defined. (* constant 720 *) Definition l_e_st_eq_landau_n_satz10g : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_not (l_e_st_eq_landau_n_lessis x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_or_th3 (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) (l_ec3e23 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10b x y) m) (l_ec3e21 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10b x y) m)))). Time Defined. (* constant 721 *) Definition l_e_st_eq_landau_n_satz10h : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_not (l_e_st_eq_landau_n_moreis x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_or_th3 (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_ec3e32 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10b x y) l) (l_ec3e31 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10b x y) l)))). Time Defined. (* constant 722 *) Definition l_e_st_eq_landau_n_satz10j : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (n:l_not (l_e_st_eq_landau_n_moreis x y)), l_e_st_eq_landau_n_less x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (n:l_not (l_e_st_eq_landau_n_moreis x y)) => l_or3e3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10a x y) (l_or_th5 (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) n) (l_or_th4 (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) n)))). Time Defined. (* constant 723 *) Definition l_e_st_eq_landau_n_satz10k : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (n:l_not (l_e_st_eq_landau_n_lessis x y)), l_e_st_eq_landau_n_more x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (n:l_not (l_e_st_eq_landau_n_lessis x y)) => l_or3e2 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_satz10a x y) (l_or_th4 (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) n) (l_or_th5 (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) n)))). Time Defined. (* constant 724 *) Definition l_e_st_eq_landau_n_315_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less y z), (forall (v:l_e_st_eq_landau_n_nat), (forall (dv:l_e_st_eq_landau_n_diffprop y x v), (forall (w:l_e_st_eq_landau_n_nat), (forall (dw:l_e_st_eq_landau_n_diffprop z y w), l_e_st_eq_landau_n_is z (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl v w))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less y z) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => (fun (w:l_e_st_eq_landau_n_nat) => (fun (dw:l_e_st_eq_landau_n_diffprop z y w) => l_e_tr3is l_e_st_eq_landau_n_nat z (l_e_st_eq_landau_n_pl y w) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl x v) w) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_pl v w)) dw (l_e_st_eq_landau_n_ispl1 y (l_e_st_eq_landau_n_pl x v) w dv) (l_e_st_eq_landau_n_asspl1 x v w)))))))))). Time Defined. (* constant 725 *) Definition l_e_st_eq_landau_n_315_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less y z), (forall (v:l_e_st_eq_landau_n_nat), (forall (dv:l_e_st_eq_landau_n_diffprop y x v), (forall (w:l_e_st_eq_landau_n_nat), (forall (dw:l_e_st_eq_landau_n_diffprop z y w), l_e_st_eq_landau_n_less x z))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less y z) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => (fun (w:l_e_st_eq_landau_n_nat) => (fun (dw:l_e_st_eq_landau_n_diffprop z y w) => l_somei l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop z x u) (l_e_st_eq_landau_n_pl v w) (l_e_st_eq_landau_n_315_t1 x y z l k v dv w dw)))))))))). Time Defined. (* constant 726 *) Definition l_e_st_eq_landau_n_315_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less y z), (forall (v:l_e_st_eq_landau_n_nat), (forall (dv:l_e_st_eq_landau_n_diffprop y x v), l_e_st_eq_landau_n_less x z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less y z) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => l_someapp l_e_st_eq_landau_n_nat (fun (w:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop z y w) k (l_e_st_eq_landau_n_less x z) (fun (w:l_e_st_eq_landau_n_nat) => (fun (dw:l_e_st_eq_landau_n_diffprop z y w) => l_e_st_eq_landau_n_315_t2 x y z l k v dv w dw))))))))). Time Defined. (* constant 727 *) Definition l_e_st_eq_landau_n_satz15 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less y z), l_e_st_eq_landau_n_less x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less y z) => l_someapp l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop y x v) l (l_e_st_eq_landau_n_less x z) (fun (v:l_e_st_eq_landau_n_nat) => (fun (dv:l_e_st_eq_landau_n_diffprop y x v) => l_e_st_eq_landau_n_315_t3 x y z l k v dv))))))). Time Defined. (* constant 728 *) Definition l_e_st_eq_landau_n_trless : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less y z), l_e_st_eq_landau_n_less x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less y z) => l_e_st_eq_landau_n_satz15 x y z l k))))). Time Defined. (* constant 729 *) Definition l_e_st_eq_landau_n_trmore : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more y z), l_e_st_eq_landau_n_more x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more y z) => l_e_st_eq_landau_n_satz15 z y x n m))))). Time Defined. (* constant 730 *) Definition l_e_st_eq_landau_n_315_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more y z), l_e_st_eq_landau_n_more x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more y z) => l_e_st_eq_landau_n_satz12 z x (l_e_st_eq_landau_n_satz15 z y x (l_e_st_eq_landau_n_satz11 y z n) (l_e_st_eq_landau_n_satz11 x y m))))))). Time Defined. (* constant 731 *) Definition l_e_st_eq_landau_n_satz16a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_less y z), l_e_st_eq_landau_n_less x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_less y z) => l_orapp (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_less x z) l (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_trless x y z u k) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_isless1 y x z (l_e_symis l_e_st_eq_landau_n_nat x y u) k)))))). Time Defined. (* constant 732 *) Definition l_e_st_eq_landau_n_satz16b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_lessis y z), l_e_st_eq_landau_n_less x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => l_orapp (l_e_st_eq_landau_n_less y z) (l_e_st_eq_landau_n_is y z) (l_e_st_eq_landau_n_less x z) k (fun (u:l_e_st_eq_landau_n_less y z) => l_e_st_eq_landau_n_trless x y z l u) (fun (u:l_e_st_eq_landau_n_is y z) => l_e_st_eq_landau_n_isless2 y z x u l)))))). Time Defined. (* constant 733 *) Definition l_e_st_eq_landau_n_satz16c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_more y z), l_e_st_eq_landau_n_more x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_more y z) => l_e_st_eq_landau_n_satz16b z y x n (l_e_st_eq_landau_n_satz13 x y m)))))). Time Defined. (* constant 734 *) Definition l_e_st_eq_landau_n_satz16d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_moreis y z), l_e_st_eq_landau_n_more x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_moreis y z) => l_e_st_eq_landau_n_satz16a z y x (l_e_st_eq_landau_n_satz13 y z n) m))))). Time Defined. (* constant 735 *) Definition l_e_st_eq_landau_n_317_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is y z), l_e_st_eq_landau_n_lessis x z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is y z) => l_e_st_eq_landau_n_lessisi2 x z (l_e_tris l_e_st_eq_landau_n_nat x y z i j)))))))). Time Defined. (* constant 736 *) Definition l_e_st_eq_landau_n_317_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_less y z), l_e_st_eq_landau_n_lessis x z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_less y z) => l_e_st_eq_landau_n_lessisi1 x z (l_e_st_eq_landau_n_satz16a x y z l j)))))))). Time Defined. (* constant 737 *) Definition l_e_st_eq_landau_n_317_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_lessis x z)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => (fun (i:l_e_st_eq_landau_n_is x y) => l_orapp (l_e_st_eq_landau_n_less y z) (l_e_st_eq_landau_n_is y z) (l_e_st_eq_landau_n_lessis x z) k (fun (u:l_e_st_eq_landau_n_less y z) => l_e_st_eq_landau_n_317_t2 x y z l k i u) (fun (u:l_e_st_eq_landau_n_is y z) => l_e_st_eq_landau_n_317_t1 x y z l k i u))))))). Time Defined. (* constant 738 *) Definition l_e_st_eq_landau_n_317_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), (forall (j:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_lessis x z)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => (fun (j:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_lessisi1 x z (l_e_st_eq_landau_n_satz16b x y z j k))))))). Time Defined. (* constant 739 *) Definition l_e_st_eq_landau_n_satz17 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), l_e_st_eq_landau_n_lessis x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => l_orapp (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_lessis x z) l (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_317_t4 x y z l k u) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_317_t3 x y z l k u)))))). Time Defined. (* constant 740 *) Definition l_e_st_eq_landau_n_317_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), (forall (j:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_lessis x z)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => (fun (j:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_lessisi1 x z (l_e_st_eq_landau_n_satz16b x y z j k))))))). Time Defined. (* constant 741 *) Definition l_e_st_eq_landau_n_317_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_lessis x z)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_islessis1 y x z (l_e_symis l_e_st_eq_landau_n_nat x y i) k)))))). Time Defined. (* constant 742 *) Definition l_e_st_eq_landau_n_317_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), l_e_st_eq_landau_n_lessis x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => l_orapp (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_lessis x z) l (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_317_t5 x y z l k u) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_317_t6 x y z l k u)))))). Time Defined. (* constant 743 *) Definition l_e_st_eq_landau_n_trlessis : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis y z), l_e_st_eq_landau_n_lessis x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis y z) => l_e_st_eq_landau_n_satz17 x y z l k))))). Time Defined. (* constant 744 *) Definition l_e_st_eq_landau_n_trmoreis : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis y z), l_e_st_eq_landau_n_moreis x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis y z) => l_e_st_eq_landau_n_satz14 z x (l_e_st_eq_landau_n_satz17 z y x (l_e_st_eq_landau_n_satz13 y z n) (l_e_st_eq_landau_n_satz13 x y m))))))). Time Defined. (* constant 745 *) Definition l_e_st_eq_landau_n_satz18 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x y) x)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_somei l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_pl x y) x u) y (l_e_refis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x y)))). Time Defined. (* constant 746 *) Definition l_e_st_eq_landau_n_satz18a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_less x (l_e_st_eq_landau_n_pl x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz18 x y)). Time Defined. (* constant 747 *) Definition l_e_st_eq_landau_n_satz18b : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_suc x) x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) x (l_e_st_eq_landau_n_satz4a x) (l_e_st_eq_landau_n_satz18 x l_e_st_eq_landau_n_1)). Time Defined. (* constant 748 *) Definition l_e_st_eq_landau_n_satz18c : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_less x (l_e_st_eq_landau_n_suc x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz18b x). Time Defined. (* constant 749 *) Definition l_e_st_eq_landau_n_319_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl u y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_e_tris l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl u y) du (l_e_st_eq_landau_n_compl y u))))))). Time Defined. (* constant 750 *) Definition l_e_st_eq_landau_n_319_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl y z) u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl u y) z) (l_e_st_eq_landau_n_pl u (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl y z) u) (l_e_st_eq_landau_n_ispl1 x (l_e_st_eq_landau_n_pl u y) z (l_e_st_eq_landau_n_319_t1 x y z m u du)) (l_e_st_eq_landau_n_asspl1 u y z) (l_e_st_eq_landau_n_compl u (l_e_st_eq_landau_n_pl y z)))))))). Time Defined. (* constant 751 *) Definition l_e_st_eq_landau_n_319_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_somei l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z) v) u (l_e_st_eq_landau_n_319_t2 x y z m u du))))))). Time Defined. (* constant 752 *) Definition l_e_st_eq_landau_n_satz19a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u) m (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_diffprop x y u) => l_e_st_eq_landau_n_319_t3 x y z m u v)))))). Time Defined. (* constant 753 *) Definition l_e_st_eq_landau_n_satz19b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_ispl1 x y z i)))). Time Defined. (* constant 754 *) Definition l_e_st_eq_landau_n_satz19c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz11 (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_satz19a y x z (l_e_st_eq_landau_n_satz12 x y l)))))). Time Defined. (* constant 755 *) Definition l_e_st_eq_landau_n_319_anders1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz19a y x z l)))). Time Defined. (* constant 756 *) Definition l_e_st_eq_landau_n_satz19d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_compl x z) (l_e_st_eq_landau_n_compl y z) (l_e_st_eq_landau_n_satz19a x y z m))))). Time Defined. (* constant 757 *) Definition l_e_st_eq_landau_n_satz19e : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_ispl2 x y z i)))). Time Defined. (* constant 758 *) Definition l_e_st_eq_landau_n_satz19f : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_compl x z) (l_e_st_eq_landau_n_compl y z) (l_e_st_eq_landau_n_satz19c x y z l))))). Time Defined. (* constant 759 *) Definition l_e_st_eq_landau_n_319_anders2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz19d y x z l)))). Time Defined. (* constant 760 *) Definition l_e_st_eq_landau_n_satz19g : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (m:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (m:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_ismore2 (l_e_st_eq_landau_n_pl x u) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_ispl1 x y u i) (l_e_st_eq_landau_n_satz19d z u x m))))))). Time Defined. (* constant 761 *) Definition l_e_st_eq_landau_n_satz19h : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (m:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl u y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (m:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl u y) (l_e_st_eq_landau_n_compl x z) (l_e_st_eq_landau_n_compl y u) (l_e_st_eq_landau_n_satz19g x y z u i m))))))). Time Defined. (* constant 762 *) Definition l_e_st_eq_landau_n_satz19j : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (l:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (l:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_isless2 (l_e_st_eq_landau_n_pl x u) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_ispl1 x y u i) (l_e_st_eq_landau_n_satz19f z u x l))))))). Time Defined. (* constant 763 *) Definition l_e_st_eq_landau_n_satz19k : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (l:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl u y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (l:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl u y) (l_e_st_eq_landau_n_compl x z) (l_e_st_eq_landau_n_compl y u) (l_e_st_eq_landau_n_satz19j x y z u i l))))))). Time Defined. (* constant 764 *) Definition l_e_st_eq_landau_n_319_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_moreisi1 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_satz19a x y z n)))))). Time Defined. (* constant 765 *) Definition l_e_st_eq_landau_n_319_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_moreisi2 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_ispl1 x y z i)))))). Time Defined. (* constant 766 *) Definition l_e_st_eq_landau_n_satz19l : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => l_orapp (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) m (fun (u:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_319_t4 x y z m u) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_319_t5 x y z m u))))). Time Defined. (* constant 767 *) Definition l_e_st_eq_landau_n_satz19m : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => l_e_st_eq_landau_n_ismoreis12 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_compl x z) (l_e_st_eq_landau_n_compl y z) (l_e_st_eq_landau_n_satz19l x y z m))))). Time Defined. (* constant 768 *) Definition l_e_st_eq_landau_n_satz19n : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => l_e_st_eq_landau_n_satz13 (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_satz19l y x z (l_e_st_eq_landau_n_satz14 x y l)))))). Time Defined. (* constant 769 *) Definition l_e_st_eq_landau_n_satz19o : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => l_e_st_eq_landau_n_satz13 (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_satz19m y x z (l_e_st_eq_landau_n_satz14 x y l)))))). Time Defined. (* constant 770 *) Definition l_e_st_eq_landau_n_320_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_or3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz10a x y))). Time Defined. (* constant 771 *) Definition l_e_st_eq_landau_n_320_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_ec3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)))). Time Defined. (* constant 772 *) Definition l_e_st_eq_landau_n_satz20a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)), l_e_st_eq_landau_n_more x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) => l_ec3_th11 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_320_t1 x y z) (l_e_st_eq_landau_n_320_t2 x y z) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_satz19b x y z u) (fun (u:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz19a x y z u) (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz19c x y z u) m)))). Time Defined. (* constant 773 *) Definition l_e_st_eq_landau_n_satz20b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)), l_e_st_eq_landau_n_is x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) => l_ec3_th10 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_320_t1 x y z) (l_e_st_eq_landau_n_320_t2 x y z) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_satz19b x y z u) (fun (u:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz19a x y z u) (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz19c x y z u) i)))). Time Defined. (* constant 774 *) Definition l_e_st_eq_landau_n_satz20c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)), l_e_st_eq_landau_n_less x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) => l_ec3_th12 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_320_t1 x y z) (l_e_st_eq_landau_n_320_t2 x y z) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_satz19b x y z u) (fun (u:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz19a x y z u) (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz19c x y z u) l)))). Time Defined. (* constant 775 *) Definition l_e_st_eq_landau_n_320_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_compl z x) i (l_e_st_eq_landau_n_compl y z))))). Time Defined. (* constant 776 *) Definition l_e_st_eq_landau_n_320_andersb : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)), l_e_st_eq_landau_n_is x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) => l_e_st_eq_landau_n_satz8a z x y (l_e_st_eq_landau_n_320_t3 x y z i))))). Time Defined. (* constant 777 *) Definition l_e_st_eq_landau_n_320_andersc : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)), l_e_st_eq_landau_n_less x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z)) => l_e_st_eq_landau_n_satz20a y x z l)))). Time Defined. (* constant 778 *) Definition l_e_st_eq_landau_n_satz20d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y)), l_e_st_eq_landau_n_more x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y)) => l_e_st_eq_landau_n_satz20a x y z (l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_compl z x) (l_e_st_eq_landau_n_compl z y) m))))). Time Defined. (* constant 779 *) Definition l_e_st_eq_landau_n_satz20e : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y)), l_e_st_eq_landau_n_is x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y)) => l_e_st_eq_landau_n_satz20b x y z (l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_compl x z) i (l_e_st_eq_landau_n_compl z y)))))). Time Defined. (* constant 780 *) Definition l_e_st_eq_landau_n_satz20f : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y)), l_e_st_eq_landau_n_less x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl z y)) => l_e_st_eq_landau_n_satz20c x y z (l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_pl z x) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_compl z x) (l_e_st_eq_landau_n_compl z y) l))))). Time Defined. (* constant 781 *) Definition l_e_st_eq_landau_n_321_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_satz19a x y z m)))))). Time Defined. (* constant 782 *) Definition l_e_st_eq_landau_n_321_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_pl z y) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl u y) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_compl z y) (l_e_st_eq_landau_n_compl u y) (l_e_st_eq_landau_n_satz19a z u y n))))))). Time Defined. (* constant 783 *) Definition l_e_st_eq_landau_n_satz21 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_trmore (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_321_t1 x y z u m n) (l_e_st_eq_landau_n_321_t2 x y z u m n))))))). Time Defined. (* constant 784 *) Definition l_e_st_eq_landau_n_321_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_trmore (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_satz19a x y z m) (l_e_st_eq_landau_n_satz19d z u y n))))))). Time Defined. (* constant 785 *) Definition l_e_st_eq_landau_n_satz21a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_satz21 y x u z l k)))))). Time Defined. (* constant 786 *) Definition l_e_st_eq_landau_n_321_andersa : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_satz11 (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_satz21 y x u z (l_e_st_eq_landau_n_satz12 x y l) (l_e_st_eq_landau_n_satz12 z u k)))))))). Time Defined. (* constant 787 *) Definition l_e_st_eq_landau_n_satz22a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_orapp (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)) m (fun (v:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz21 x y z u v n) (fun (v:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_satz19g x y z u v n))))))). Time Defined. (* constant 788 *) Definition l_e_st_eq_landau_n_satz22b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_moreis z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => l_orapp (l_e_st_eq_landau_n_more z u) (l_e_st_eq_landau_n_is z u) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)) n (fun (v:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_satz21 x y z u m v) (fun (v:l_e_st_eq_landau_n_is z u) => l_e_st_eq_landau_n_satz19h z u x y v m))))))). Time Defined. (* constant 789 *) Definition l_e_st_eq_landau_n_satz22c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_satz22a y x u z (l_e_st_eq_landau_n_satz14 x y l) k)))))). Time Defined. (* constant 790 *) Definition l_e_st_eq_landau_n_satz22d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_lessis z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_lessis z u) => l_e_st_eq_landau_n_satz22b y x u z l (l_e_st_eq_landau_n_satz14 z u k))))))). Time Defined. (* constant 791 *) Definition l_e_st_eq_landau_n_323_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => l_e_st_eq_landau_n_moreisi2 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u) (l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y z) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_ispl1 x y z i) (l_e_st_eq_landau_n_ispl2 z u y j)))))))))). Time Defined. (* constant 792 *) Definition l_e_st_eq_landau_n_323_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (i:l_e_st_eq_landau_n_is x y), (forall (o:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (o:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_moreisi1 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_satz22a x y z u m o))))))))). Time Defined. (* constant 793 *) Definition l_e_st_eq_landau_n_323_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (i:l_e_st_eq_landau_n_is x y) => l_orapp (l_e_st_eq_landau_n_more z u) (l_e_st_eq_landau_n_is z u) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)) n (fun (v:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_323_t2 x y z u m n i v) (fun (v:l_e_st_eq_landau_n_is z u) => l_e_st_eq_landau_n_323_t1 x y z u m n i v)))))))). Time Defined. (* constant 794 *) Definition l_e_st_eq_landau_n_323_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (o:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (o:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_moreisi1 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_satz22b x y z u o n)))))))). Time Defined. (* constant 795 *) Definition l_e_st_eq_landau_n_satz23 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => l_orapp (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)) m (fun (v:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_323_t4 x y z u m n v) (fun (v:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_323_t3 x y z u m n v))))))). Time Defined. (* constant 796 *) Definition l_e_st_eq_landau_n_323_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (o:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (o:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_moreisi1 (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_satz22b x y z u o n)))))))). Time Defined. (* constant 797 *) Definition l_e_st_eq_landau_n_323_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_ismoreis2 (l_e_st_eq_landau_n_pl x u) (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_ispl1 x y u i) (l_e_st_eq_landau_n_satz19m z u x n)))))))). Time Defined. (* constant 798 *) Definition l_e_st_eq_landau_n_323_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => l_orapp (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u)) m (fun (v:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_323_t5 x y z u m n v) (fun (v:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_323_t6 x y z u m n v))))))). Time Defined. (* constant 799 *) Definition l_e_st_eq_landau_n_satz23a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis z u), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis z u) => l_e_st_eq_landau_n_satz13 (l_e_st_eq_landau_n_pl y u) (l_e_st_eq_landau_n_pl x z) (l_e_st_eq_landau_n_satz23 y x u z (l_e_st_eq_landau_n_satz14 x y l) (l_e_st_eq_landau_n_satz14 z u k)))))))). Time Defined. (* constant 800 *) Definition l_e_st_eq_landau_n_324_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 u))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) => l_e_tris l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_suc u) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 u) i (l_e_st_eq_landau_n_satz4g u))))). Time Defined. (* constant 801 *) Definition l_e_st_eq_landau_n_324_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)), l_e_st_eq_landau_n_more x l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) => l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 u) x l_e_st_eq_landau_n_1 (l_e_symis l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 u) (l_e_st_eq_landau_n_324_t1 x n u i)) (l_e_st_eq_landau_n_satz18 l_e_st_eq_landau_n_1 u))))). Time Defined. (* constant 802 *) Definition l_e_st_eq_landau_n_324_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_more x l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_satz3 x n) (l_e_st_eq_landau_n_more x l_e_st_eq_landau_n_1) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_suc u)) => l_e_st_eq_landau_n_324_t2 x n u v)))). Time Defined. (* constant 803 *) Definition l_e_st_eq_landau_n_satz24 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_moreis x l_e_st_eq_landau_n_1). exact (fun (x:l_e_st_eq_landau_n_nat) => l_or_th2 (l_e_st_eq_landau_n_more x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (fun (u:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_324_t3 x u)). Time Defined. (* constant 804 *) Definition l_e_st_eq_landau_n_satz24a : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz13 x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24 x)). Time Defined. (* constant 805 *) Definition l_e_st_eq_landau_n_satz24b : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_suc x) l_e_st_eq_landau_n_1). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_324_t3 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_ax3 x)). Time Defined. (* constant 806 *) Definition l_e_st_eq_landau_n_satz24c : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_suc x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz24b x). Time Defined. (* constant 807 *) Definition l_e_st_eq_landau_n_325_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more y x), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop y x u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_pl x u) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more y x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop y x u) => l_e_st_eq_landau_n_satz19m u l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_satz24 u)))))). Time Defined. (* constant 808 *) Definition l_e_st_eq_landau_n_325_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more y x), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop y x u), l_e_st_eq_landau_n_moreis y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more y x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop y x u) => l_e_st_eq_landau_n_ismoreis1 (l_e_st_eq_landau_n_pl x u) y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_symis l_e_st_eq_landau_n_nat y (l_e_st_eq_landau_n_pl x u) du) (l_e_st_eq_landau_n_325_t1 x y m u du)))))). Time Defined. (* constant 809 *) Definition l_e_st_eq_landau_n_satz25 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more y x), l_e_st_eq_landau_n_moreis y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more y x) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop y x u) m (l_e_st_eq_landau_n_moreis y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_diffprop y x u) => l_e_st_eq_landau_n_325_t2 x y m u v))))). Time Defined. (* constant 810 *) Definition l_e_st_eq_landau_n_satz25a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more y x), l_e_st_eq_landau_n_moreis y (l_e_st_eq_landau_n_suc x)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more y x) => l_e_st_eq_landau_n_ismoreis2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) y (l_e_st_eq_landau_n_satz4a x) (l_e_st_eq_landau_n_satz25 x y m)))). Time Defined. (* constant 811 *) Definition l_e_st_eq_landau_n_satz25b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less y x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less y x) => l_e_st_eq_landau_n_satz13 x (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz25 y x l)))). Time Defined. (* constant 812 *) Definition l_e_st_eq_landau_n_satz25c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less y x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc y) x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less y x) => l_e_st_eq_landau_n_islessis1 (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc y) x (l_e_st_eq_landau_n_satz4a y) (l_e_st_eq_landau_n_satz25b x y l)))). Time Defined. (* constant 813 *) Definition l_e_st_eq_landau_n_326_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (m:l_e_st_eq_landau_n_more y x), l_e_st_eq_landau_n_moreis y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (m:l_e_st_eq_landau_n_more y x) => l_e_st_eq_landau_n_satz25 x y m)))). Time Defined. (* constant 814 *) Definition l_e_st_eq_landau_n_326_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_not (l_e_st_eq_landau_n_more y x)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => l_imp_th3 (l_e_st_eq_landau_n_more y x) (l_e_st_eq_landau_n_moreis y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_satz10h y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) l) (fun (v:l_e_st_eq_landau_n_more y x) => l_e_st_eq_landau_n_326_t1 x y l v)))). Time Defined. (* constant 815 *) Definition l_e_st_eq_landau_n_satz26 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_lessis y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less y (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_satz10e y x (l_e_st_eq_landau_n_326_t2 x y l)))). Time Defined. (* constant 816 *) Definition l_e_st_eq_landau_n_satz26a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less y (l_e_st_eq_landau_n_suc x)), l_e_st_eq_landau_n_lessis y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less y (l_e_st_eq_landau_n_suc x)) => l_e_st_eq_landau_n_satz26 x y (l_e_st_eq_landau_n_isless2 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) y (l_e_st_eq_landau_n_satz4e x) l)))). Time Defined. (* constant 817 *) Definition l_e_st_eq_landau_n_satz26b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) x), l_e_st_eq_landau_n_moreis y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) x) => l_e_st_eq_landau_n_satz14 x y (l_e_st_eq_landau_n_satz26 y x m)))). Time Defined. (* constant 818 *) Definition l_e_st_eq_landau_n_satz26c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_suc y) x), l_e_st_eq_landau_n_moreis y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_suc y) x) => l_e_st_eq_landau_n_satz26b x y (l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_satz4e y) m)))). Time Defined. (* constant 819 *) Definition l_e_st_eq_landau_n_327_lbprop : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), Prop))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => l_imp (p m) (l_e_st_eq_landau_n_lessis n m)))). Time Defined. (* constant 820 *) Definition l_e_st_eq_landau_n_lb : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), Prop)). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_all (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_327_lbprop p n x))). Time Defined. (* constant 821 *) Definition l_e_st_eq_landau_n_min : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), Prop)). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p n) (p n))). Time Defined. (* constant 822 *) Definition l_e_st_eq_landau_n_327_t1 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_327_lbprop p l_e_st_eq_landau_n_1 n))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (x:p n) => l_e_st_eq_landau_n_satz24a n)))). Time Defined. (* constant 823 *) Definition l_e_st_eq_landau_n_327_t2 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), l_e_st_eq_landau_n_lb p l_e_st_eq_landau_n_1)). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_327_t1 p s x))). Time Defined. (* constant 824 *) Definition l_e_st_eq_landau_n_327_t3 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)), (forall (y:l_e_st_eq_landau_n_nat), (forall (yp:p y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) y))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (yp:p y) => l_e_st_eq_landau_n_satz18 y l_e_st_eq_landau_n_1))))). Time Defined. (* constant 825 *) Definition l_e_st_eq_landau_n_327_t4 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)), (forall (y:l_e_st_eq_landau_n_nat), (forall (yp:p y), l_not (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) y)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (yp:p y) => l_e_st_eq_landau_n_satz10g (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) y (l_e_st_eq_landau_n_327_t3 p s l y yp)))))). Time Defined. (* constant 826 *) Definition l_e_st_eq_landau_n_327_t5 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)), (forall (y:l_e_st_eq_landau_n_nat), (forall (yp:p y), l_not (l_e_st_eq_landau_n_327_lbprop p (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) y)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (yp:p y) => l_imp_th4 (p y) (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) y) yp (l_e_st_eq_landau_n_327_t4 p s l y yp)))))). Time Defined. (* constant 827 *) Definition l_e_st_eq_landau_n_327_t6 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)), (forall (y:l_e_st_eq_landau_n_nat), (forall (yp:p y), l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (yp:p y) => l_all_th1 l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_327_lbprop p (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) x) y (l_e_st_eq_landau_n_327_t5 p s l y yp)))))). Time Defined. (* constant 828 *) Definition l_e_st_eq_landau_n_327_t7 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)), (forall (y:l_e_st_eq_landau_n_nat), (forall (yp:p y), l_con))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (yp:p y) => l_mp (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1)) l_con (l (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_327_t6 p s l y yp)))))). Time Defined. (* constant 829 *) Definition l_e_st_eq_landau_n_327_t8 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)), l_con))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (l:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)) => l_someapp l_e_st_eq_landau_n_nat p s l_con (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:p x) => l_e_st_eq_landau_n_327_t7 p s l x y))))). Time Defined. (* constant 830 *) Definition l_e_st_eq_landau_n_327_t9 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))), (forall (m:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p m), l_not (l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1))))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p m) => n m))))). Time Defined. (* constant 831 *) Definition l_e_st_eq_landau_n_327_t10 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))), (forall (m:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p m), l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p m) => l_et (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)) (l_and_th3 (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_327_t9 p s n m l) l)))))). Time Defined. (* constant 832 *) Definition l_e_st_eq_landau_n_327_t11 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))), (forall (m:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p m), l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_suc m)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p m) => l_e_isp l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lb p x) (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc m) (l_e_st_eq_landau_n_327_t10 p s n m l) (l_e_st_eq_landau_n_satz4a m)))))). Time Defined. (* constant 833 *) Definition l_e_st_eq_landau_n_327_t12 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p x)))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lb p y) (l_e_st_eq_landau_n_327_t2 p s) (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_lb p y) => l_e_st_eq_landau_n_327_t11 p s n y z)) x)))). Time Defined. (* constant 834 *) Definition l_e_st_eq_landau_n_327_t13 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (x:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) => l_e_st_eq_landau_n_327_t8 p s (l_e_st_eq_landau_n_327_t12 p s x)))). Time Defined. (* constant 835 *) Definition l_e_st_eq_landau_n_327_t14 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_lb p m)))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => l_ande1 (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1))) a)))). Time Defined. (* constant 836 *) Definition l_e_st_eq_landau_n_327_t15 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => l_ande2 (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1))) a)))). Time Defined. (* constant 837 *) Definition l_e_st_eq_landau_n_327_t16 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), (forall (nmp:l_not (p m)), (forall (n:l_e_st_eq_landau_n_nat), (forall (np:p n), l_e_st_eq_landau_n_lessis m n))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => (fun (nmp:l_not (p m)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (np:p n) => l_mp (p n) (l_e_st_eq_landau_n_lessis m n) np (l_e_st_eq_landau_n_327_t14 p s m a n)))))))). Time Defined. (* constant 838 *) Definition l_e_st_eq_landau_n_327_t17 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), (forall (nmp:l_not (p m)), (forall (n:l_e_st_eq_landau_n_nat), (forall (np:p n), l_not (l_e_st_eq_landau_n_is m n)))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => (fun (nmp:l_not (p m)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (np:p n) => l_imp_th3 (l_e_st_eq_landau_n_is m n) (p m) nmp (fun (x:l_e_st_eq_landau_n_is m n) => l_e_isp l_e_st_eq_landau_n_nat p n m np (l_e_symis l_e_st_eq_landau_n_nat m n x))))))))). Time Defined. (* constant 839 *) Definition l_e_st_eq_landau_n_327_t18 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), (forall (nmp:l_not (p m)), (forall (n:l_e_st_eq_landau_n_nat), (forall (np:p n), l_e_st_eq_landau_n_less m n))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => (fun (nmp:l_not (p m)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (np:p n) => l_ore1 (l_e_st_eq_landau_n_less m n) (l_e_st_eq_landau_n_is m n) (l_e_st_eq_landau_n_327_t16 p s m a nmp n np) (l_e_st_eq_landau_n_327_t17 p s m a nmp n np)))))))). Time Defined. (* constant 840 *) Definition l_e_st_eq_landau_n_327_t19 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), (forall (nmp:l_not (p m)), (forall (n:l_e_st_eq_landau_n_nat), (forall (np:p n), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1) n))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => (fun (nmp:l_not (p m)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (np:p n) => l_e_st_eq_landau_n_satz25b n m (l_e_st_eq_landau_n_327_t18 p s m a nmp n np)))))))). Time Defined. (* constant 841 *) Definition l_e_st_eq_landau_n_327_t20 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), (forall (nmp:l_not (p m)), l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => (fun (nmp:l_not (p m)) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:p x) => l_e_st_eq_landau_n_327_t19 p s m a nmp x y))))))). Time Defined. (* constant 842 *) Definition l_e_st_eq_landau_n_327_t21 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), (forall (nmp:l_not (p m)), l_con))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => (fun (nmp:l_not (p m)) => l_mp (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)) l_con (l_e_st_eq_landau_n_327_t20 p s m a nmp) (l_e_st_eq_landau_n_327_t15 p s m a)))))). Time Defined. (* constant 843 *) Definition l_e_st_eq_landau_n_327_t22 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), p m)))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => l_et (p m) (fun (x:l_not (p m)) => l_e_st_eq_landau_n_327_t21 p s m a x))))). Time Defined. (* constant 844 *) Definition l_e_st_eq_landau_n_327_t23 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_min p m)))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (a:l_and (l_e_st_eq_landau_n_lb p m) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl m l_e_st_eq_landau_n_1)))) => l_andi (l_e_st_eq_landau_n_lb p m) (p m) (l_e_st_eq_landau_n_327_t14 p s m a) (l_e_st_eq_landau_n_327_t22 p s m a))))). Time Defined. (* constant 845 *) Definition l_e_st_eq_landau_n_satz27 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => l_some_th6 l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x) (l_e_st_eq_landau_n_327_t13 p s) (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_and (l_e_st_eq_landau_n_lb p x) (l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_st_eq_landau_n_327_t23 p s x y)))). Time Defined. (* constant 846 *) Definition l_e_st_eq_landau_n_327_t24 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_327_lbprop p l_e_st_eq_landau_n_1 u))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (x:p u) => l_e_st_eq_landau_n_satz24a u)))). Time Defined. (* constant 847 *) Definition l_e_st_eq_landau_n_327_t25 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), l_e_st_eq_landau_n_lb p l_e_st_eq_landau_n_1)). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_327_t24 p n x))). Time Defined. (* constant 848 *) Definition l_e_st_eq_landau_n_327_t26 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p u), l_not (l_e_st_eq_landau_n_min p u))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p u) => n u)))). Time Defined. (* constant 849 *) Definition l_e_st_eq_landau_n_327_t27 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p u), l_not (p u))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p u) => l_and_th3 (l_e_st_eq_landau_n_lb p u) (p u) (l_e_st_eq_landau_n_327_t26 p n u l) l)))). Time Defined. (* constant 850 *) Definition l_e_st_eq_landau_n_327_t28 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p u), (forall (v:l_e_st_eq_landau_n_nat), (forall (vp:p v), l_e_st_eq_landau_n_nis u v)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p u) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (vp:p v) => l_imp_th3 (l_e_st_eq_landau_n_is u v) (p u) (l_e_st_eq_landau_n_327_t27 p n u l) (fun (x:l_e_st_eq_landau_n_is u v) => l_e_isp1 l_e_st_eq_landau_n_nat p v u vp x))))))). Time Defined. (* constant 851 *) Definition l_e_st_eq_landau_n_327_t29 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p u), (forall (v:l_e_st_eq_landau_n_nat), (forall (vp:p v), l_e_st_eq_landau_n_lessis u v)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p u) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (vp:p v) => l_mp (p v) (l_e_st_eq_landau_n_lessis u v) vp (l v))))))). Time Defined. (* constant 852 *) Definition l_e_st_eq_landau_n_327_t30 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p u), (forall (v:l_e_st_eq_landau_n_nat), (forall (vp:p v), l_e_st_eq_landau_n_less u v)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p u) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (vp:p v) => l_ore1 (l_e_st_eq_landau_n_less u v) (l_e_st_eq_landau_n_is u v) (l_e_st_eq_landau_n_327_t29 p n u l v vp) (l_e_st_eq_landau_n_327_t28 p n u l v vp))))))). Time Defined. (* constant 853 *) Definition l_e_st_eq_landau_n_327_t31 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p u), (forall (v:l_e_st_eq_landau_n_nat), (forall (vp:p v), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) v)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p u) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (vp:p v) => l_e_st_eq_landau_n_satz25c v u (l_e_st_eq_landau_n_327_t30 p n u l v vp))))))). Time Defined. (* constant 854 *) Definition l_e_st_eq_landau_n_327_t32 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p u), (forall (v:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_327_lbprop p (l_e_st_eq_landau_n_suc u) v))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p u) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (x:p v) => l_e_st_eq_landau_n_327_t31 p n u l v x)))))). Time Defined. (* constant 855 *) Definition l_e_st_eq_landau_n_327_t33 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lb p u), l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_suc u))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lb p u) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_327_t32 p n u l x))))). Time Defined. (* constant 856 *) Definition l_e_st_eq_landau_n_327_t34 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lb p u))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lb p x) (l_e_st_eq_landau_n_327_t25 p n) (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_lb p x) => l_e_st_eq_landau_n_327_t33 p n x y)) u))). Time Defined. (* constant 857 *) Definition l_e_st_eq_landau_n_327_t35 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (u:l_e_st_eq_landau_n_nat), (forall (up:p u), l_not (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) u))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (up:p u) => l_e_st_eq_landau_n_satz10g (l_e_st_eq_landau_n_suc u) u (l_e_st_eq_landau_n_satz18b u))))). Time Defined. (* constant 858 *) Definition l_e_st_eq_landau_n_327_t36 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (u:l_e_st_eq_landau_n_nat), (forall (up:p u), l_not (l_e_st_eq_landau_n_327_lbprop p (l_e_st_eq_landau_n_suc u) u))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (up:p u) => l_imp_th4 (p u) (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) u) up (l_e_st_eq_landau_n_327_t35 p s u up))))). Time Defined. (* constant 859 *) Definition l_e_st_eq_landau_n_327_t37 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (u:l_e_st_eq_landau_n_nat), (forall (up:p u), l_not (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_suc u)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (up:p u) => l_all_th1 l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_327_lbprop p (l_e_st_eq_landau_n_suc u) x) u (l_e_st_eq_landau_n_327_t36 p s u up))))). Time Defined. (* constant 860 *) Definition l_e_st_eq_landau_n_327_t38 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), (forall (u:l_e_st_eq_landau_n_nat), (forall (up:p u), l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (up:p u) => (fun (y:l_none l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) => l_mp (l_e_st_eq_landau_n_lb p (l_e_st_eq_landau_n_suc u)) l_con (l_e_st_eq_landau_n_327_t34 p y (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_327_t37 p s u up)))))). Time Defined. (* constant 861 *) Definition l_e_st_eq_landau_n_327_anders : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => l_someapp l_e_st_eq_landau_n_nat p s (l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)) (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:p x) => l_e_st_eq_landau_n_327_t38 p s x y)))). Time Defined. (* constant 862 *) Definition l_e_st_eq_landau_n_327_t39 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), l_e_st_eq_landau_n_lb p n))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_ande1 (l_e_st_eq_landau_n_lb p n) (p n) mn))))). Time Defined. (* constant 863 *) Definition l_e_st_eq_landau_n_327_t40 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), l_e_st_eq_landau_n_lb p m))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_ande1 (l_e_st_eq_landau_n_lb p m) (p m) mm))))). Time Defined. (* constant 864 *) Definition l_e_st_eq_landau_n_327_t41 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), p n))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_ande2 (l_e_st_eq_landau_n_lb p n) (p n) mn))))). Time Defined. (* constant 865 *) Definition l_e_st_eq_landau_n_327_t42 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), p m))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_ande2 (l_e_st_eq_landau_n_lb p m) (p m) mm))))). Time Defined. (* constant 866 *) Definition l_e_st_eq_landau_n_327_t43 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), l_e_st_eq_landau_n_327_lbprop p n m))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_e_st_eq_landau_n_327_t39 p n m mn mm m))))). Time Defined. (* constant 867 *) Definition l_e_st_eq_landau_n_327_t44 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), l_e_st_eq_landau_n_327_lbprop p m n))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_e_st_eq_landau_n_327_t40 p n m mn mm n))))). Time Defined. (* constant 868 *) Definition l_e_st_eq_landau_n_327_t45 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), l_e_st_eq_landau_n_lessis n m))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_mp (p m) (l_e_st_eq_landau_n_lessis n m) (l_e_st_eq_landau_n_327_t42 p n m mn mm) (l_e_st_eq_landau_n_327_t43 p n m mn mm)))))). Time Defined. (* constant 869 *) Definition l_e_st_eq_landau_n_327_t46 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), l_e_st_eq_landau_n_lessis m n))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_mp (p n) (l_e_st_eq_landau_n_lessis m n) (l_e_st_eq_landau_n_327_t41 p n m mn mm) (l_e_st_eq_landau_n_327_t44 p n m mn mm)))))). Time Defined. (* constant 870 *) Definition l_e_st_eq_landau_n_327_t47 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (mn:l_e_st_eq_landau_n_min p n), (forall (mm:l_e_st_eq_landau_n_min p m), l_e_st_eq_landau_n_is n m))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mn:l_e_st_eq_landau_n_min p n) => (fun (mm:l_e_st_eq_landau_n_min p m) => l_ore2 (l_e_st_eq_landau_n_more n m) (l_e_st_eq_landau_n_is n m) (l_e_st_eq_landau_n_satz14 m n (l_e_st_eq_landau_n_327_t46 p n m mn mm)) (l_e_st_eq_landau_n_satz10d n m (l_e_st_eq_landau_n_327_t45 p n m mn mm))))))). Time Defined. (* constant 871 *) Definition l_e_st_eq_landau_n_327_t48 : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), l_e_amone l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x)). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_min p x) => (fun (v:l_e_st_eq_landau_n_min p y) => l_e_st_eq_landau_n_327_t47 p x y u v))))). Time Defined. (* constant 872 *) Definition l_e_st_eq_landau_n_satz27a : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), (forall (s:l_e_st_eq_landau_n_some p), l_e_st_eq_landau_n_one (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => (fun (s:l_e_st_eq_landau_n_some p) => l_e_onei l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min p x) (l_e_st_eq_landau_n_327_t48 p) (l_e_st_eq_landau_n_satz27 p s))). Time Defined. (* constant 873 *) Definition l_e_st_eq_landau_n_428_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_all (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (f y) x)))). Time Defined. (* constant 874 *) Definition l_e_st_eq_landau_n_428_prop2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_and (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x f))). Time Defined. (* constant 875 *) Definition l_e_st_eq_landau_n_428_prop3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), Prop)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (a y) (b y))))))). Time Defined. (* constant 876 *) Definition l_e_st_eq_landau_n_428_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), l_e_st_eq_landau_n_is (a l_e_st_eq_landau_n_1) x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => l_ande1 (l_e_st_eq_landau_n_is (a l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x a) pa))))). Time Defined. (* constant 877 *) Definition l_e_st_eq_landau_n_428_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), l_e_st_eq_landau_n_is (b l_e_st_eq_landau_n_1) x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => l_ande1 (l_e_st_eq_landau_n_is (b l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x b) pb))))). Time Defined. (* constant 878 *) Definition l_e_st_eq_landau_n_428_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), l_e_st_eq_landau_n_428_prop3 x a b pa pb l_e_st_eq_landau_n_1))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => l_e_tris2 l_e_st_eq_landau_n_nat (a l_e_st_eq_landau_n_1) (b l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_428_t1 x a b pa pb) (l_e_st_eq_landau_n_428_t2 x a b pa pb)))))). Time Defined. (* constant 879 *) Definition l_e_st_eq_landau_n_428_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (a y) x) (l_e_st_eq_landau_n_pl (b y) x)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y) => l_e_st_eq_landau_n_ispl1 (a y) (b y) x p))))))). Time Defined. (* constant 880 *) Definition l_e_st_eq_landau_n_428_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y), l_e_st_eq_landau_n_428_prop1 x a))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y) => l_ande2 (l_e_st_eq_landau_n_is (a l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x a) pa))))))). Time Defined. (* constant 881 *) Definition l_e_st_eq_landau_n_428_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y), l_e_st_eq_landau_n_428_prop1 x b))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y) => l_ande2 (l_e_st_eq_landau_n_is (b l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x b) pb))))))). Time Defined. (* constant 882 *) Definition l_e_st_eq_landau_n_428_t7 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y), l_e_st_eq_landau_n_is (a (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (a y) x)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y) => l_e_st_eq_landau_n_428_t5 x a b pa pb y p y))))))). Time Defined. (* constant 883 *) Definition l_e_st_eq_landau_n_428_t8 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y), l_e_st_eq_landau_n_is (b (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (b y) x)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y) => l_e_st_eq_landau_n_428_t6 x a b pa pb y p y))))))). Time Defined. (* constant 884 *) Definition l_e_st_eq_landau_n_428_t9 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y), l_e_st_eq_landau_n_428_prop3 x a b pa pb (l_e_st_eq_landau_n_suc y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop3 x a b pa pb y) => l_e_tr3is l_e_st_eq_landau_n_nat (a (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (a y) x) (l_e_st_eq_landau_n_pl (b y) x) (b (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_428_t7 x a b pa pb y p) (l_e_st_eq_landau_n_428_t4 x a b pa pb y p) (l_e_symis l_e_st_eq_landau_n_nat (b (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (b y) x) (l_e_st_eq_landau_n_428_t8 x a b pa pb y p))))))))). Time Defined. (* constant 885 *) Definition l_e_st_eq_landau_n_428_t10 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_428_prop3 x a b pa pb y)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_428_prop3 x a b pa pb z) (l_e_st_eq_landau_n_428_t3 x a b pa pb) (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_428_prop3 x a b pa pb z) => l_e_st_eq_landau_n_428_t9 x a b pa pb z u)) y)))))). Time Defined. (* constant 886 *) Definition l_e_st_eq_landau_n_428_t11 : (forall (x:l_e_st_eq_landau_n_nat), (forall (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pa:l_e_st_eq_landau_n_428_prop2 x a), (forall (pb:l_e_st_eq_landau_n_428_prop2 x b), l_e_is (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) a b))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (a:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (b:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pa:l_e_st_eq_landau_n_428_prop2 x a) => (fun (pb:l_e_st_eq_landau_n_428_prop2 x b) => l_e_fisi l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat a b (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_428_t10 x a b pa pb y)))))). Time Defined. (* constant 887 *) Definition l_e_st_eq_landau_n_428_a1 : (forall (x:l_e_st_eq_landau_n_nat), l_e_amone (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_428_prop2 x z)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (u:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (v:l_e_st_eq_landau_n_428_prop2 x z) => (fun (w:l_e_st_eq_landau_n_428_prop2 x u) => l_e_st_eq_landau_n_428_t11 x z u v w))))). Time Defined. (* constant 888 *) Definition l_e_st_eq_landau_n_428_prop4 : (forall (x:l_e_st_eq_landau_n_nat), Prop). exact (fun (x:l_e_st_eq_landau_n_nat) => l_some (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_428_prop2 x z)). Time Defined. (* constant 889 *) Definition l_e_st_eq_landau_n_428_id : (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat). exact (fun (y:l_e_st_eq_landau_n_nat) => y). Time Defined. (* constant 890 *) Definition l_e_st_eq_landau_n_428_t12 : l_e_st_eq_landau_n_428_prop1 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_428_id. exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz4e x). Time Defined. (* constant 891 *) Definition l_e_st_eq_landau_n_428_t13 : l_e_st_eq_landau_n_428_prop2 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_428_id. exact (l_andi (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_428_id l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_428_prop1 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_428_id) (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_428_t12). Time Defined. (* constant 892 *) Definition l_e_st_eq_landau_n_428_t14 : l_e_st_eq_landau_n_428_prop4 l_e_st_eq_landau_n_1. exact (l_somei (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_428_prop2 l_e_st_eq_landau_n_1 z) l_e_st_eq_landau_n_428_id l_e_st_eq_landau_n_428_t13). Time Defined. (* constant 893 *) Definition l_e_st_eq_landau_n_428_g : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl (f y) y))))). Time Defined. (* constant 894 *) Definition l_e_st_eq_landau_n_428_t15 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) x)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => l_ande1 (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x f) pf)))). Time Defined. (* constant 895 *) Definition l_e_st_eq_landau_n_428_t16 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_428_g x p f pf l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_428_g x p f pf l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_ispl1 (f l_e_st_eq_landau_n_1) x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_428_t15 x p f pf)) (l_e_st_eq_landau_n_satz4a x))))). Time Defined. (* constant 896 *) Definition l_e_st_eq_landau_n_428_t17 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_428_prop1 x f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_ande2 (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x f) pf))))). Time Defined. (* constant 897 *) Definition l_e_st_eq_landau_n_428_t18 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (f y) x)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_428_t17 x p f pf y y))))). Time Defined. (* constant 898 *) Definition l_e_st_eq_landau_n_428_t19 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_428_g x p f pf (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (f y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_428_g x p f pf (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (f y) x) (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (f y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y))) (l_e_st_eq_landau_n_ispl1 (f (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (f y) x) (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_428_t18 x p f pf y)) (l_e_st_eq_landau_n_asspl1 (f y) x (l_e_st_eq_landau_n_suc y))))))). Time Defined. (* constant 899 *) Definition l_e_st_eq_landau_n_428_t20 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc x))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_satz4b x y) (l_e_st_eq_landau_n_satz4h x y) (l_e_st_eq_landau_n_compl (l_e_st_eq_landau_n_suc x) y)))))). Time Defined. (* constant 900 *) Definition l_e_st_eq_landau_n_428_t21 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_428_g x p f pf (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_428_g x p f pf y) (l_e_st_eq_landau_n_suc x))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_428_g x p f pf (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (f y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y))) (l_e_st_eq_landau_n_pl (f y) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc x))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_428_g x p f pf y) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_428_t19 x p f pf y) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc x)) (f y) (l_e_st_eq_landau_n_428_t20 x p f pf y)) (l_e_st_eq_landau_n_asspl2 (f y) y (l_e_st_eq_landau_n_suc x))))))). Time Defined. (* constant 901 *) Definition l_e_st_eq_landau_n_428_t22 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), l_e_st_eq_landau_n_428_prop1 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_428_g x p f pf))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_428_t21 x p f pf y))))). Time Defined. (* constant 902 *) Definition l_e_st_eq_landau_n_428_t23 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), l_e_st_eq_landau_n_428_prop2 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_428_g x p f pf))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => l_andi (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_428_g x p f pf l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_428_prop1 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_428_g x p f pf)) (l_e_st_eq_landau_n_428_t16 x p f pf) (l_e_st_eq_landau_n_428_t22 x p f pf))))). Time Defined. (* constant 903 *) Definition l_e_st_eq_landau_n_428_t24 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), (forall (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (pf:l_e_st_eq_landau_n_428_prop2 x f), l_e_st_eq_landau_n_428_prop4 (l_e_st_eq_landau_n_suc x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => (fun (f:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (pf:l_e_st_eq_landau_n_428_prop2 x f) => l_somei (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_428_prop2 (l_e_st_eq_landau_n_suc x) z) (l_e_st_eq_landau_n_428_g x p f pf) (l_e_st_eq_landau_n_428_t23 x p f pf))))). Time Defined. (* constant 904 *) Definition l_e_st_eq_landau_n_428_t25 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_428_prop4 x), l_e_st_eq_landau_n_428_prop4 (l_e_st_eq_landau_n_suc x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_428_prop4 x) => l_someapp (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_428_prop2 x z) p (l_e_st_eq_landau_n_428_prop4 (l_e_st_eq_landau_n_suc x)) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (u:l_e_st_eq_landau_n_428_prop2 x z) => l_e_st_eq_landau_n_428_t24 x p z u)))). Time Defined. (* constant 905 *) Definition l_e_st_eq_landau_n_428_b1 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_428_prop4 x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_428_prop4 y) l_e_st_eq_landau_n_428_t14 (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_428_prop4 y) => l_e_st_eq_landau_n_428_t25 y u)) x). Time Defined. (* constant 906 *) Definition l_e_st_eq_landau_n_satz28 : (forall (x:l_e_st_eq_landau_n_nat), l_e_one (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_and (l_e_st_eq_landau_n_is (z l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_all (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (z (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (z y) x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_onei (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_428_prop2 x z) (l_e_st_eq_landau_n_428_a1 x) (l_e_st_eq_landau_n_428_b1 x)). Time Defined. (* constant 907 *) Definition l_e_st_eq_landau_n_times : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_ind (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_428_prop2 x z) (l_e_st_eq_landau_n_satz28 x)). Time Defined. (* constant 908 *) Definition l_e_st_eq_landau_n_ts : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_times x y)). Time Defined. (* constant 909 *) Definition l_e_st_eq_landau_n_428_t26 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_428_prop2 x (l_e_st_eq_landau_n_times x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_oneax (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (z:(forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_428_prop2 x z) (l_e_st_eq_landau_n_satz28 x)). Time Defined. (* constant 910 *) Definition l_e_st_eq_landau_n_satz28a : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x l_e_st_eq_landau_n_1) x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_ande1 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_times x l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x (l_e_st_eq_landau_n_times x)) (l_e_st_eq_landau_n_428_t26 x)). Time Defined. (* constant 911 *) Definition l_e_st_eq_landau_n_428_t27 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_428_prop1 x (l_e_st_eq_landau_n_times x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_ande2 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_times x l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_428_prop1 x (l_e_st_eq_landau_n_times x)) (l_e_st_eq_landau_n_428_t26 x)). Time Defined. (* constant 912 *) Definition l_e_st_eq_landau_n_satz28b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_428_t27 x y)). Time Defined. (* constant 913 *) Definition l_e_st_eq_landau_n_428_t28 : l_e_is (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (l_e_st_eq_landau_n_times l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_428_id. exact (l_e_st_eq_landau_n_428_t11 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_times l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_428_id (l_e_st_eq_landau_n_428_t26 l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_428_t13). Time Defined. (* constant 914 *) Definition l_e_st_eq_landau_n_satz28c : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 x) x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_fise l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_times l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_428_id l_e_st_eq_landau_n_428_t28 x). Time Defined. (* constant 915 *) Definition l_e_st_eq_landau_n_428_t29 : (forall (x:l_e_st_eq_landau_n_nat), l_e_is (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (l_e_st_eq_landau_n_times (l_e_st_eq_landau_n_suc x)) (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_times x y) y)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_428_t11 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_times (l_e_st_eq_landau_n_suc x)) (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_times x y) y) (l_e_st_eq_landau_n_428_t26 (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_428_t23 x (l_e_st_eq_landau_n_428_b1 x) (l_e_st_eq_landau_n_times x) (l_e_st_eq_landau_n_428_t26 x))). Time Defined. (* constant 916 *) Definition l_e_st_eq_landau_n_satz28d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_fise l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_times (l_e_st_eq_landau_n_suc x)) (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_times x z) z) (l_e_st_eq_landau_n_428_t29 x) y)). Time Defined. (* constant 917 *) Definition l_e_st_eq_landau_n_satz28e : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_ts x l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_satz28a x)). Time Defined. (* constant 918 *) Definition l_e_st_eq_landau_n_satz28f : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) x) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_suc y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) x) (l_e_st_eq_landau_n_satz28b x y))). Time Defined. (* constant 919 *) Definition l_e_st_eq_landau_n_satz28g : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 x) x (l_e_st_eq_landau_n_satz28c x)). Time Defined. (* constant 920 *) Definition l_e_st_eq_landau_n_satz28h : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) y) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc x) y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) y) (l_e_st_eq_landau_n_satz28d x y))). Time Defined. (* constant 921 *) Definition l_e_st_eq_landau_n_ists1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ts u z) x y i)))). Time Defined. (* constant 922 *) Definition l_e_st_eq_landau_n_ists2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ts z u) x y i)))). Time Defined. (* constant 923 *) Definition l_e_st_eq_landau_n_ists12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ists1 x y z i) (l_e_st_eq_landau_n_ists2 z u y j))))))). Time Defined. (* constant 924 *) Definition l_e_st_eq_landau_n_429_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts y x))). Time Defined. (* constant 925 *) Definition l_e_st_eq_landau_n_429_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts y l_e_st_eq_landau_n_1) y)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz28a y)). Time Defined. (* constant 926 *) Definition l_e_st_eq_landau_n_429_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 y) y)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz28c y)). Time Defined. (* constant 927 *) Definition l_e_st_eq_landau_n_429_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_429_prop1 l_e_st_eq_landau_n_1 y)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris2 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 y) (l_e_st_eq_landau_n_ts y l_e_st_eq_landau_n_1) y (l_e_st_eq_landau_n_429_t2 x y) (l_e_st_eq_landau_n_429_t1 x y))). Time Defined. (* constant 928 *) Definition l_e_st_eq_landau_n_429_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_429_prop1 x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) y) (l_e_st_eq_landau_n_ts y (l_e_st_eq_landau_n_suc x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_429_prop1 x y) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts y x) y) (l_e_st_eq_landau_n_ts y (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts y x) y p) (l_e_st_eq_landau_n_satz28f y x)))). Time Defined. (* constant 929 *) Definition l_e_st_eq_landau_n_429_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_429_prop1 x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_429_prop1 x y) => l_e_st_eq_landau_n_satz28d x y))). Time Defined. (* constant 930 *) Definition l_e_st_eq_landau_n_429_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_429_prop1 x y), l_e_st_eq_landau_n_429_prop1 (l_e_st_eq_landau_n_suc x) y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_429_prop1 x y) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc x) y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) y) (l_e_st_eq_landau_n_ts y (l_e_st_eq_landau_n_suc x)) (l_e_st_eq_landau_n_429_t5 x y p) (l_e_st_eq_landau_n_429_t4 x y p)))). Time Defined. (* constant 931 *) Definition l_e_st_eq_landau_n_satz29 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_429_prop1 z y) (l_e_st_eq_landau_n_429_t3 x y) (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_429_prop1 z y) => l_e_st_eq_landau_n_429_t6 z y u)) x)). Time Defined. (* constant 932 *) Definition l_e_st_eq_landau_n_comts : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz29 x y)). Time Defined. (* constant 933 *) Definition l_e_st_eq_landau_n_429_t7 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_429_prop1 x l_e_st_eq_landau_n_1). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 x) (l_e_st_eq_landau_n_satz28a x) (l_e_st_eq_landau_n_satz28g x)). Time Defined. (* constant 934 *) Definition l_e_st_eq_landau_n_429_t8 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_429_prop1 x y), l_e_st_eq_landau_n_429_prop1 x (l_e_st_eq_landau_n_suc y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_429_prop1 x y) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts y x) x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc y) x) (l_e_st_eq_landau_n_satz28b x y) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts y x) x p) (l_e_st_eq_landau_n_satz28h y x)))). Time Defined. (* constant 935 *) Definition l_e_st_eq_landau_n_429_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts y x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_429_prop1 x z) (l_e_st_eq_landau_n_429_t7 x) (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_429_prop1 x z) => l_e_st_eq_landau_n_429_t8 x z u)) y)). Time Defined. (* constant 936 *) Definition l_e_st_eq_landau_n_430_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), Prop))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z))))). Time Defined. (* constant 937 *) Definition l_e_st_eq_landau_n_430_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_430_prop1 x y l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_suc y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc y) x (l_e_st_eq_landau_n_satz4a y)) (l_e_st_eq_landau_n_satz28b x y) (l_e_st_eq_landau_n_ispl2 x (l_e_st_eq_landau_n_ts x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_satz28e x)))). Time Defined. (* constant 938 *) Definition l_e_st_eq_landau_n_430_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_430_prop1 x y z), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z)) x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_430_prop1 x y z) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc z))) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl y z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y z)) x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z)) x) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc z)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_pl y z)) x (l_e_st_eq_landau_n_satz4b y z)) (l_e_st_eq_landau_n_satz28b x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z)) x p))))). Time Defined. (* constant 939 *) Definition l_e_st_eq_landau_n_430_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_430_prop1 x y z), l_e_st_eq_landau_n_430_prop1 x y (l_e_st_eq_landau_n_suc z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_430_prop1 x y z) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_suc z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z)) x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x z) x)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_suc z))) (l_e_st_eq_landau_n_430_t2 x y z p) (l_e_st_eq_landau_n_asspl1 (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z) x) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x z) x) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_suc z)) (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_satz28f x z)))))). Time Defined. (* constant 940 *) Definition l_e_st_eq_landau_n_satz30 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_430_prop1 x y u) (l_e_st_eq_landau_n_430_t1 x y) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_430_prop1 x y u) => l_e_st_eq_landau_n_430_t3 x y u v)) z))). Time Defined. (* constant 941 *) Definition l_e_st_eq_landau_n_disttp1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_ts z (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts z y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_satz30 z x y) (l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts z y) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_comts z x) (l_e_st_eq_landau_n_comts z y))))). Time Defined. (* constant 942 *) Definition l_e_st_eq_landau_n_disttp2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz30 x y z))). Time Defined. (* constant 943 *) Definition l_e_st_eq_landau_n_distpt1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl x y) z)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl x y) z) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_disttp1 x y z)))). Time Defined. (* constant 944 *) Definition l_e_st_eq_landau_n_distpt2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z)) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x z)) (l_e_st_eq_landau_n_disttp2 x y z)))). Time Defined. (* constant 945 *) Definition l_e_st_eq_landau_n_431_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), Prop))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) z) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y z))))). Time Defined. (* constant 946 *) Definition l_e_st_eq_landau_n_431_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_431_prop1 x y l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_satz28a (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_ists2 y (l_e_st_eq_landau_n_ts y l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_satz28e y)))). Time Defined. (* constant 947 *) Definition l_e_st_eq_landau_n_431_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_431_prop1 x y z), l_e_st_eq_landau_n_431_prop1 x y (l_e_st_eq_landau_n_suc z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_431_prop1 x y z) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_suc z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) z) (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts y z) y)) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y (l_e_st_eq_landau_n_suc z))) (l_e_st_eq_landau_n_satz28b (l_e_st_eq_landau_n_ts x y) z) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) z) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_ts x y) p) (l_e_st_eq_landau_n_distpt2 x (l_e_st_eq_landau_n_ts y z) y) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts y z) y) (l_e_st_eq_landau_n_ts y (l_e_st_eq_landau_n_suc z)) x (l_e_st_eq_landau_n_satz28f y z)))))). Time Defined. (* constant 948 *) Definition l_e_st_eq_landau_n_satz31 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) z) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_431_prop1 x y u) (l_e_st_eq_landau_n_431_t1 x y) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_431_prop1 x y u) => l_e_st_eq_landau_n_431_t2 x y u v)) z))). Time Defined. (* constant 949 *) Definition l_e_st_eq_landau_n_assts1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) z) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz31 x y z))). Time Defined. (* constant 950 *) Definition l_e_st_eq_landau_n_assts2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) z)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x y) z) (l_e_st_eq_landau_n_ts x (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_assts1 x y z)))). Time Defined. (* constant 951 *) Definition l_e_st_eq_landau_n_432_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts u z)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl y u) z) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts u z)) (l_e_st_eq_landau_n_ists1 x (l_e_st_eq_landau_n_pl y u) z du) (l_e_st_eq_landau_n_disttp1 y u z))))))). Time Defined. (* constant 952 *) Definition l_e_st_eq_landau_n_432_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (u:l_e_st_eq_landau_n_nat), (forall (du:l_e_st_eq_landau_n_diffprop x y u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (du:l_e_st_eq_landau_n_diffprop x y u) => l_somei l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z) v) (l_e_st_eq_landau_n_ts u z) (l_e_st_eq_landau_n_432_t1 x y z m u du))))))). Time Defined. (* constant 953 *) Definition l_e_st_eq_landau_n_satz32a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y u) m (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_diffprop x y u) => l_e_st_eq_landau_n_432_t2 x y z m u v)))))). Time Defined. (* constant 954 *) Definition l_e_st_eq_landau_n_satz32b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_ists1 x y z i)))). Time Defined. (* constant 955 *) Definition l_e_st_eq_landau_n_satz32c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz11 (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_satz32a y x z (l_e_st_eq_landau_n_satz12 x y l)))))). Time Defined. (* constant 956 *) Definition l_e_st_eq_landau_n_432_anders1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz32a y x z l)))). Time Defined. (* constant 957 *) Definition l_e_st_eq_landau_n_satz32d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts z y) (l_e_st_eq_landau_n_comts x z) (l_e_st_eq_landau_n_comts y z) (l_e_st_eq_landau_n_satz32a x y z m))))). Time Defined. (* constant 958 *) Definition l_e_st_eq_landau_n_satz32e : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_ists2 x y z i)))). Time Defined. (* constant 959 *) Definition l_e_st_eq_landau_n_satz32f : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts z y) (l_e_st_eq_landau_n_comts x z) (l_e_st_eq_landau_n_comts y z) (l_e_st_eq_landau_n_satz32c x y z l))))). Time Defined. (* constant 960 *) Definition l_e_st_eq_landau_n_432_anders2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz32d y x z l)))). Time Defined. (* constant 961 *) Definition l_e_st_eq_landau_n_satz32g : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (m:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (m:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_ismore2 (l_e_st_eq_landau_n_ts x u) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ists1 x y u i) (l_e_st_eq_landau_n_satz32d z u x m))))))). Time Defined. (* constant 962 *) Definition l_e_st_eq_landau_n_satz32h : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (m:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts u y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (m:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ts u y) (l_e_st_eq_landau_n_comts x z) (l_e_st_eq_landau_n_comts y u) (l_e_st_eq_landau_n_satz32g x y z u i m))))))). Time Defined. (* constant 963 *) Definition l_e_st_eq_landau_n_satz32j : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (l:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (l:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_isless2 (l_e_st_eq_landau_n_ts x u) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ists1 x y u i) (l_e_st_eq_landau_n_satz32f z u x l))))))). Time Defined. (* constant 964 *) Definition l_e_st_eq_landau_n_satz32k : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), (forall (l:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts u y))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (l:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ts u y) (l_e_st_eq_landau_n_comts x z) (l_e_st_eq_landau_n_comts y u) (l_e_st_eq_landau_n_satz32j x y z u i l))))))). Time Defined. (* constant 965 *) Definition l_e_st_eq_landau_n_432_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_moreisi1 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_satz32a x y z n)))))). Time Defined. (* constant 966 *) Definition l_e_st_eq_landau_n_432_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_moreisi2 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ists1 x y z i)))))). Time Defined. (* constant 967 *) Definition l_e_st_eq_landau_n_satz32l : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => l_orapp (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) m (fun (u:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_432_t3 x y z m u) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_432_t4 x y z m u))))). Time Defined. (* constant 968 *) Definition l_e_st_eq_landau_n_satz32m : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => l_e_st_eq_landau_n_ismoreis12 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts z y) (l_e_st_eq_landau_n_comts x z) (l_e_st_eq_landau_n_comts y z) (l_e_st_eq_landau_n_satz32l x y z m))))). Time Defined. (* constant 969 *) Definition l_e_st_eq_landau_n_satz32n : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => l_e_st_eq_landau_n_satz13 (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_satz32l y x z (l_e_st_eq_landau_n_satz14 x y l)))))). Time Defined. (* constant 970 *) Definition l_e_st_eq_landau_n_satz32o : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_ts z y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => l_e_st_eq_landau_n_satz13 (l_e_st_eq_landau_n_ts z y) (l_e_st_eq_landau_n_ts z x) (l_e_st_eq_landau_n_satz32m y x z (l_e_st_eq_landau_n_satz14 x y l)))))). Time Defined. (* constant 971 *) Definition l_e_st_eq_landau_n_433_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_or3 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz10a x y))). Time Defined. (* constant 972 *) Definition l_e_st_eq_landau_n_433_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), l_ec3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)))). Time Defined. (* constant 973 *) Definition l_e_st_eq_landau_n_satz33a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)), l_e_st_eq_landau_n_more x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) => l_ec3_th11 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_433_t1 x y z) (l_e_st_eq_landau_n_433_t2 x y z) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_satz32b x y z u) (fun (u:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz32a x y z u) (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz32c x y z u) m)))). Time Defined. (* constant 974 *) Definition l_e_st_eq_landau_n_satz33b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)), l_e_st_eq_landau_n_is x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) => l_ec3_th10 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_433_t1 x y z) (l_e_st_eq_landau_n_433_t2 x y z) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_satz32b x y z u) (fun (u:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz32a x y z u) (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz32c x y z u) i)))). Time Defined. (* constant 975 *) Definition l_e_st_eq_landau_n_satz33c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)), l_e_st_eq_landau_n_less x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) => l_ec3_th12 (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_less x y) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) (l_e_st_eq_landau_n_433_t1 x y z) (l_e_st_eq_landau_n_433_t2 x y z) (fun (u:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_satz32b x y z u) (fun (u:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz32a x y z u) (fun (u:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_satz32c x y z u) l)))). Time Defined. (* constant 976 *) Definition l_e_st_eq_landau_n_433_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)), l_e_st_eq_landau_n_less x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z)) => l_e_st_eq_landau_n_satz33a y x z l)))). Time Defined. (* constant 977 *) Definition l_e_st_eq_landau_n_434_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_satz32a x y z m)))))). Time Defined. (* constant 978 *) Definition l_e_st_eq_landau_n_434_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts z y) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts u y) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_comts z y) (l_e_st_eq_landau_n_comts u y) (l_e_st_eq_landau_n_satz32a z u y n))))))). Time Defined. (* constant 979 *) Definition l_e_st_eq_landau_n_satz34 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_trmore (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_434_t1 x y z u m n) (l_e_st_eq_landau_n_434_t2 x y z u m n))))))). Time Defined. (* constant 980 *) Definition l_e_st_eq_landau_n_434_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_trmore (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_satz32a x y z m) (l_e_st_eq_landau_n_satz32d z u y n))))))). Time Defined. (* constant 981 *) Definition l_e_st_eq_landau_n_satz34a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_satz34 y x u z l k)))))). Time Defined. (* constant 982 *) Definition l_e_st_eq_landau_n_434_andersa : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_satz11 (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_satz34 y x u z (l_e_st_eq_landau_n_satz12 x y l) (l_e_st_eq_landau_n_satz12 z u k)))))))). Time Defined. (* constant 983 *) Definition l_e_st_eq_landau_n_satz35a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_more z u) => l_orapp (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)) m (fun (v:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz34 x y z u v n) (fun (v:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_satz32g x y z u v n))))))). Time Defined. (* constant 984 *) Definition l_e_st_eq_landau_n_satz35b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (n:l_e_st_eq_landau_n_moreis z u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => l_orapp (l_e_st_eq_landau_n_more z u) (l_e_st_eq_landau_n_is z u) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)) n (fun (v:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_satz34 x y z u m v) (fun (v:l_e_st_eq_landau_n_is z u) => l_e_st_eq_landau_n_satz32h z u x y v m))))))). Time Defined. (* constant 985 *) Definition l_e_st_eq_landau_n_satz35c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_less z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_less z u) => l_e_st_eq_landau_n_satz35a y x u z (l_e_st_eq_landau_n_satz14 x y l) k)))))). Time Defined. (* constant 986 *) Definition l_e_st_eq_landau_n_satz35d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (k:l_e_st_eq_landau_n_lessis z u), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (k:l_e_st_eq_landau_n_lessis z u) => l_e_st_eq_landau_n_satz35b y x u z l (l_e_st_eq_landau_n_satz14 z u k))))))). Time Defined. (* constant 987 *) Definition l_e_st_eq_landau_n_436_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => l_e_st_eq_landau_n_moreisi2 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u) (l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y z) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ists1 x y z i) (l_e_st_eq_landau_n_ists2 z u y j)))))))))). Time Defined. (* constant 988 *) Definition l_e_st_eq_landau_n_436_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (i:l_e_st_eq_landau_n_is x y), (forall (o:l_e_st_eq_landau_n_more z u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (o:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_moreisi1 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_satz35a x y z u m o))))))))). Time Defined. (* constant 989 *) Definition l_e_st_eq_landau_n_436_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (i:l_e_st_eq_landau_n_is x y) => l_orapp (l_e_st_eq_landau_n_more z u) (l_e_st_eq_landau_n_is z u) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)) n (fun (v:l_e_st_eq_landau_n_more z u) => l_e_st_eq_landau_n_436_t2 x y z u m n i v) (fun (v:l_e_st_eq_landau_n_is z u) => l_e_st_eq_landau_n_436_t1 x y z u m n i v)))))))). Time Defined. (* constant 990 *) Definition l_e_st_eq_landau_n_436_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (o:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (o:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_moreisi1 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_satz35b x y z u o n)))))))). Time Defined. (* constant 991 *) Definition l_e_st_eq_landau_n_satz36 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => l_orapp (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)) m (fun (v:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_436_t4 x y z u m n v) (fun (v:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_436_t3 x y z u m n v))))))). Time Defined. (* constant 992 *) Definition l_e_st_eq_landau_n_436_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (o:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (o:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_moreisi1 (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_satz35b x y z u o n)))))))). Time Defined. (* constant 993 *) Definition l_e_st_eq_landau_n_436_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_ismoreis2 (l_e_st_eq_landau_n_ts x u) (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ists1 x y u i) (l_e_st_eq_landau_n_satz32m z u x n)))))))). Time Defined. (* constant 994 *) Definition l_e_st_eq_landau_n_436_anders : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moreis x y), (forall (n:l_e_st_eq_landau_n_moreis z u), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moreis x y) => (fun (n:l_e_st_eq_landau_n_moreis z u) => l_orapp (l_e_st_eq_landau_n_more x y) (l_e_st_eq_landau_n_is x y) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u)) m (fun (v:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_436_t5 x y z u m n v) (fun (v:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_436_t6 x y z u m n v))))))). Time Defined. (* constant 995 *) Definition l_e_st_eq_landau_n_satz36a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x y), (forall (k:l_e_st_eq_landau_n_lessis z u), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x y) => (fun (k:l_e_st_eq_landau_n_lessis z u) => l_e_st_eq_landau_n_satz13 (l_e_st_eq_landau_n_ts y u) (l_e_st_eq_landau_n_ts x z) (l_e_st_eq_landau_n_satz36 y x u z (l_e_st_eq_landau_n_satz14 x y l) (l_e_st_eq_landau_n_satz14 z u k)))))))). Time Defined. (* constant 996 *) Definition l_e_st_eq_landau_n_mn_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_one (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y z)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_onei l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y z) (l_e_st_eq_landau_n_satz8b x y) m))). Time Defined. (* constant 997 *) Definition l_e_st_eq_landau_n_mn : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_nat))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_ind l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y z) (l_e_st_eq_landau_n_mn_t1 x y m)))). Time Defined. (* constant 998 *) Definition l_e_st_eq_landau_n_mn_th1a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_mn x y m))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_oneax l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x y z) (l_e_st_eq_landau_n_mn_t1 x y m)))). Time Defined. (* constant 999 *) Definition l_e_st_eq_landau_n_mn_th1b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_mn x y m)) x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_symis l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_mn x y m)) (l_e_st_eq_landau_n_mn_th1a x y m)))). Time Defined. (* constant 1000 *) Definition l_e_st_eq_landau_n_mn_th1c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_mn x y m) y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_tris l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl y (l_e_st_eq_landau_n_mn x y m)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_mn x y m) y) (l_e_st_eq_landau_n_mn_th1a x y m) (l_e_st_eq_landau_n_compl y (l_e_st_eq_landau_n_mn x y m))))). Time Defined. (* constant 1001 *) Definition l_e_st_eq_landau_n_mn_th1d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_mn x y m) y) x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_symis l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_mn x y m) y) (l_e_st_eq_landau_n_mn_th1c x y m)))). Time Defined. (* constant 1002 *) Definition l_e_st_eq_landau_n_mn_th1e : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl y z) x), l_e_st_eq_landau_n_is z (l_e_st_eq_landau_n_mn x y m)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl y z) x) => l_e_st_eq_landau_n_satz8b x y z (l_e_st_eq_landau_n_mn x y m) (l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl y z) x i) (l_e_st_eq_landau_n_mn_th1a x y m)))))). Time Defined. (* constant 1003 *) Definition l_e_st_eq_landau_n_mn_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x z), (forall (n:l_e_st_eq_landau_n_more y u), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl u (l_e_st_eq_landau_n_mn x z m)) y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x z) => (fun (n:l_e_st_eq_landau_n_more y u) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl u (l_e_st_eq_landau_n_mn x z m)) (l_e_st_eq_landau_n_pl z (l_e_st_eq_landau_n_mn x z m)) x y (l_e_st_eq_landau_n_ispl1 u z (l_e_st_eq_landau_n_mn x z m) (l_e_symis l_e_st_eq_landau_n_nat z u j)) (l_e_st_eq_landau_n_mn_th1b x z m) i)))))))). Time Defined. (* constant 1004 *) Definition l_e_st_eq_landau_n_ismn12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (z:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x z), (forall (n:l_e_st_eq_landau_n_more y u), (forall (i:l_e_st_eq_landau_n_is x y), (forall (j:l_e_st_eq_landau_n_is z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_mn x z m) (l_e_st_eq_landau_n_mn y u n))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x z) => (fun (n:l_e_st_eq_landau_n_more y u) => (fun (i:l_e_st_eq_landau_n_is x y) => (fun (j:l_e_st_eq_landau_n_is z u) => l_e_st_eq_landau_n_mn_th1e y u (l_e_st_eq_landau_n_mn x z m) n (l_e_st_eq_landau_n_mn_t2 x y z u m n i j))))))))). Time Defined. (* constant 1005 *) Definition l_e_st_eq_landau_n_1to : (forall (n:l_e_st_eq_landau_n_nat), Type). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_ot l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis x n)). Time Defined. (* constant 1006 *) Definition l_e_st_eq_landau_n_outn : (forall (n:l_e_st_eq_landau_n_nat), (forall (x:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x n), l_e_st_eq_landau_n_1to n))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x n) => l_e_out l_e_st_eq_landau_n_nat (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis y n) x l))). Time Defined. (* constant 1007 *) Definition l_e_st_eq_landau_n_inn : (forall (n:l_e_st_eq_landau_n_nat), (forall (xn:l_e_st_eq_landau_n_1to n), l_e_st_eq_landau_n_nat)). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (xn:l_e_st_eq_landau_n_1to n) => l_e_in l_e_st_eq_landau_n_nat (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis y n) xn)). Time Defined. (* constant 1008 *) Definition l_e_st_eq_landau_n_1top : (forall (n:l_e_st_eq_landau_n_nat), (forall (xn:l_e_st_eq_landau_n_1to n), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_inn n xn) n)). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (xn:l_e_st_eq_landau_n_1to n) => l_e_inp l_e_st_eq_landau_n_nat (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis y n) xn)). Time Defined. (* constant 1009 *) Definition l_e_st_eq_landau_n_isoutni : (forall (n:l_e_st_eq_landau_n_nat), (forall (x:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x n), (forall (y:l_e_st_eq_landau_n_nat), (forall (k:l_e_st_eq_landau_n_lessis y n), (forall (i:l_e_st_eq_landau_n_is x y), l_e_is (l_e_st_eq_landau_n_1to n) (l_e_st_eq_landau_n_outn n x l) (l_e_st_eq_landau_n_outn n y k))))))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x n) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (k:l_e_st_eq_landau_n_lessis y n) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isouti l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis z n) x l y k i)))))). Time Defined. (* constant 1010 *) Definition l_e_st_eq_landau_n_isoutne : (forall (n:l_e_st_eq_landau_n_nat), (forall (x:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x n), (forall (y:l_e_st_eq_landau_n_nat), (forall (k:l_e_st_eq_landau_n_lessis y n), (forall (i:l_e_is (l_e_st_eq_landau_n_1to n) (l_e_st_eq_landau_n_outn n x l) (l_e_st_eq_landau_n_outn n y k)), l_e_st_eq_landau_n_is x y)))))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x n) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (k:l_e_st_eq_landau_n_lessis y n) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to n) (l_e_st_eq_landau_n_outn n x l) (l_e_st_eq_landau_n_outn n y k)) => l_e_isoute l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis z n) x l y k i)))))). Time Defined. (* constant 1011 *) Definition l_e_st_eq_landau_n_isinni : (forall (n:l_e_st_eq_landau_n_nat), (forall (xn:l_e_st_eq_landau_n_1to n), (forall (yn:l_e_st_eq_landau_n_1to n), (forall (i:l_e_is (l_e_st_eq_landau_n_1to n) xn yn), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_inn n xn) (l_e_st_eq_landau_n_inn n yn))))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (xn:l_e_st_eq_landau_n_1to n) => (fun (yn:l_e_st_eq_landau_n_1to n) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to n) xn yn) => l_e_isini l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis z n) xn yn i)))). Time Defined. (* constant 1012 *) Definition l_e_st_eq_landau_n_isinne : (forall (n:l_e_st_eq_landau_n_nat), (forall (xn:l_e_st_eq_landau_n_1to n), (forall (yn:l_e_st_eq_landau_n_1to n), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_inn n xn) (l_e_st_eq_landau_n_inn n yn)), l_e_is (l_e_st_eq_landau_n_1to n) xn yn)))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (xn:l_e_st_eq_landau_n_1to n) => (fun (yn:l_e_st_eq_landau_n_1to n) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_inn n xn) (l_e_st_eq_landau_n_inn n yn)) => l_e_isine l_e_st_eq_landau_n_nat (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis z n) xn yn i)))). Time Defined. (* constant 1013 *) Definition l_e_st_eq_landau_n_isoutinn : (forall (n:l_e_st_eq_landau_n_nat), (forall (xn:l_e_st_eq_landau_n_1to n), l_e_is (l_e_st_eq_landau_n_1to n) xn (l_e_st_eq_landau_n_outn n (l_e_st_eq_landau_n_inn n xn) (l_e_st_eq_landau_n_1top n xn)))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (xn:l_e_st_eq_landau_n_1to n) => l_e_isoutin l_e_st_eq_landau_n_nat (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis y n) xn)). Time Defined. (* constant 1014 *) Definition l_e_st_eq_landau_n_isinoutn : (forall (n:l_e_st_eq_landau_n_nat), (forall (x:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x n), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_inn n (l_e_st_eq_landau_n_outn n x l))))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x n) => l_e_isinout l_e_st_eq_landau_n_nat (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis y n) x l))). Time Defined. (* constant 1015 *) Definition l_e_st_eq_landau_n_1o : l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1. exact (l_e_st_eq_landau_n_outn l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_lessisi2 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1 (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1))). Time Defined. (* constant 1016 *) Definition l_e_st_eq_landau_n_singlet_u0 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_nat). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_inn l_e_st_eq_landau_n_1 u). Time Defined. (* constant 1017 *) Definition l_e_st_eq_landau_n_singlet_t1 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_singlet_u0 u) l_e_st_eq_landau_n_1). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_1top l_e_st_eq_landau_n_1 u). Time Defined. (* constant 1018 *) Definition l_e_st_eq_landau_n_singlet_t2 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_singlet_u0 u) l_e_st_eq_landau_n_1). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_ore2 (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_singlet_u0 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_singlet_u0 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz24 (l_e_st_eq_landau_n_singlet_u0 u)) (l_e_st_eq_landau_n_satz10d (l_e_st_eq_landau_n_singlet_u0 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_singlet_t1 u))). Time Defined. (* constant 1019 *) Definition l_e_st_eq_landau_n_singlet_th1 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) u l_e_st_eq_landau_n_1o). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_tris (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) u (l_e_st_eq_landau_n_outn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_singlet_u0 u) (l_e_st_eq_landau_n_singlet_t1 u)) l_e_st_eq_landau_n_1o (l_e_st_eq_landau_n_isoutinn l_e_st_eq_landau_n_1 u) (l_e_st_eq_landau_n_isoutni l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_singlet_u0 u) (l_e_st_eq_landau_n_singlet_t1 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_lessisi2 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1 (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_singlet_t2 u))). Time Defined. (* constant 1020 *) Definition l_e_st_eq_landau_n_2 : l_e_st_eq_landau_n_nat. exact (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1). Time Defined. (* constant 1021 *) Definition l_e_st_eq_landau_n_pair_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x l_e_st_eq_landau_n_2), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_lessis x l_e_st_eq_landau_n_1))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x l_e_st_eq_landau_n_2) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_2) => l_e_st_eq_landau_n_satz26 l_e_st_eq_landau_n_1 x (l_ore1 (l_e_st_eq_landau_n_less x l_e_st_eq_landau_n_2) (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_2) l n)))). Time Defined. (* constant 1022 *) Definition l_e_st_eq_landau_n_pair_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x l_e_st_eq_landau_n_2), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x l_e_st_eq_landau_n_2) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_2) => l_ore2 (l_e_st_eq_landau_n_more x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz24 x) (l_e_st_eq_landau_n_satz10d x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_pair_t1 x l n))))). Time Defined. (* constant 1023 *) Definition l_e_st_eq_landau_n_pair_th1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis x l_e_st_eq_landau_n_2), l_or (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_2))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis x l_e_st_eq_landau_n_2) => l_or_th2 (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_2) (fun (t:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_2) => l_e_st_eq_landau_n_pair_t2 x l t))). Time Defined. (* constant 1024 *) Definition l_e_st_eq_landau_n_pair_th2 : l_e_st_eq_landau_n_nis l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1. exact (l_e_notis_th1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_ax3 l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz4e l_e_st_eq_landau_n_1)). Time Defined. (* constant 1025 *) Definition l_e_st_eq_landau_n_1t : l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2. exact (l_e_st_eq_landau_n_outn l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a l_e_st_eq_landau_n_2)). Time Defined. (* constant 1026 *) Definition l_e_st_eq_landau_n_2t : l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2. exact (l_e_st_eq_landau_n_outn l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_lessisi2 l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_2 (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_2))). Time Defined. (* constant 1027 *) Definition l_e_st_eq_landau_n_pair_u0 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_nat). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => l_e_st_eq_landau_n_inn l_e_st_eq_landau_n_2 u). Time Defined. (* constant 1028 *) Definition l_e_st_eq_landau_n_pair_t3 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_2). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => l_e_st_eq_landau_n_1top l_e_st_eq_landau_n_2 u). Time Defined. (* constant 1029 *) Definition l_e_st_eq_landau_n_pair_t4 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) (l_e_st_eq_landau_n_outn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_pair_u0 u) (l_e_st_eq_landau_n_pair_t3 u)) l_e_st_eq_landau_n_1t)). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isoutni l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_pair_u0 u) (l_e_st_eq_landau_n_pair_t3 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a l_e_st_eq_landau_n_2) i)). Time Defined. (* constant 1030 *) Definition l_e_st_eq_landau_n_pair_t5 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t)). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_1) => l_e_tris (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u (l_e_st_eq_landau_n_outn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_pair_u0 u) (l_e_st_eq_landau_n_pair_t3 u)) l_e_st_eq_landau_n_1t (l_e_st_eq_landau_n_isoutinn l_e_st_eq_landau_n_2 u) (l_e_st_eq_landau_n_pair_t4 u i))). Time Defined. (* constant 1031 *) Definition l_e_st_eq_landau_n_pair_t6 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_2), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) (l_e_st_eq_landau_n_outn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_pair_u0 u) (l_e_st_eq_landau_n_pair_t3 u)) l_e_st_eq_landau_n_2t)). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_2) => l_e_st_eq_landau_n_isoutni l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_pair_u0 u) (l_e_st_eq_landau_n_pair_t3 u) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_lessisi2 l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_2 (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_2)) i)). Time Defined. (* constant 1032 *) Definition l_e_st_eq_landau_n_pair_t7 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_2), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t)). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_2) => l_e_tris (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u (l_e_st_eq_landau_n_outn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_pair_u0 u) (l_e_st_eq_landau_n_pair_t3 u)) l_e_st_eq_landau_n_2t (l_e_st_eq_landau_n_isoutinn l_e_st_eq_landau_n_2 u) (l_e_st_eq_landau_n_pair_t6 u i))). Time Defined. (* constant 1033 *) Definition l_e_st_eq_landau_n_pair_th3 : (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_or (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t) (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t)). exact (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => l_or_th9 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_2) (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t) (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t) (l_e_st_eq_landau_n_pair_th1 (l_e_st_eq_landau_n_pair_u0 u) (l_e_st_eq_landau_n_pair_t3 u)) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_pair_t5 u t) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 u) l_e_st_eq_landau_n_2) => l_e_st_eq_landau_n_pair_t7 u t)). Time Defined. (* constant 1034 *) Definition l_e_st_eq_landau_n_pair_t9 : (forall (i:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pair_u0 l_e_st_eq_landau_n_2t) (l_e_st_eq_landau_n_pair_u0 l_e_st_eq_landau_n_1t)). exact (fun (i:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t) => l_e_isini l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessis x l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t i). Time Defined. (* constant 1035 *) Definition l_e_st_eq_landau_n_pair_t10 : (forall (i:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1). exact (fun (i:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t) => l_e_tr3is l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_pair_u0 l_e_st_eq_landau_n_2t) (l_e_st_eq_landau_n_pair_u0 l_e_st_eq_landau_n_1t) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_isinoutn l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_lessisi2 l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_2 (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_2))) (l_e_st_eq_landau_n_pair_t9 i) (l_e_symis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_pair_u0 l_e_st_eq_landau_n_1t) (l_e_st_eq_landau_n_isinoutn l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a l_e_st_eq_landau_n_2)))). Time Defined. (* constant 1036 *) Definition l_e_st_eq_landau_n_pair_th4 : l_not (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t). exact (l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t) (l_e_st_eq_landau_n_is l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_pair_th2 (fun (t:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t) => l_e_st_eq_landau_n_pair_t10 t)). Time Defined. (* constant 1037 *) Definition l_e_st_eq_landau_n_pair1type : (forall (alpha:Type), Type). exact (fun (alpha:Type) => (forall (x:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), alpha)). Time Defined. (* constant 1038 *) Definition l_e_st_eq_landau_n_pair1 : (forall (alpha:Type), (forall (a:alpha), (forall (b:alpha), l_e_st_eq_landau_n_pair1type alpha))). exact (fun (alpha:Type) => (fun (a:alpha) => (fun (b:alpha) => (fun (x:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => l_e_ite (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) x l_e_st_eq_landau_n_1t) alpha a b)))). Time Defined. (* constant 1039 *) Definition l_e_st_eq_landau_n_first1 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), alpha)). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => p l_e_st_eq_landau_n_1t)). Time Defined. (* constant 1040 *) Definition l_e_st_eq_landau_n_second1 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), alpha)). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => p l_e_st_eq_landau_n_2t)). Time Defined. (* constant 1041 *) Definition l_e_st_eq_landau_n_first1is1 : (forall (alpha:Type), (forall (a:alpha), (forall (b:alpha), l_e_is alpha (l_e_st_eq_landau_n_first1 alpha (l_e_st_eq_landau_n_pair1 alpha a b)) a))). exact (fun (alpha:Type) => (fun (a:alpha) => (fun (b:alpha) => l_e_itet (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_1t l_e_st_eq_landau_n_1t) alpha a b (l_e_refis (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_1t)))). Time Defined. (* constant 1042 *) Definition l_e_st_eq_landau_n_first1is2 : (forall (alpha:Type), (forall (a:alpha), (forall (b:alpha), l_e_is alpha a (l_e_st_eq_landau_n_first1 alpha (l_e_st_eq_landau_n_pair1 alpha a b))))). exact (fun (alpha:Type) => (fun (a:alpha) => (fun (b:alpha) => l_e_symis alpha (l_e_st_eq_landau_n_first1 alpha (l_e_st_eq_landau_n_pair1 alpha a b)) a (l_e_st_eq_landau_n_first1is1 alpha a b)))). Time Defined. (* constant 1043 *) Definition l_e_st_eq_landau_n_second1is1 : (forall (alpha:Type), (forall (a:alpha), (forall (b:alpha), l_e_is alpha (l_e_st_eq_landau_n_second1 alpha (l_e_st_eq_landau_n_pair1 alpha a b)) b))). exact (fun (alpha:Type) => (fun (a:alpha) => (fun (b:alpha) => l_e_itef (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_2t l_e_st_eq_landau_n_1t) alpha a b l_e_st_eq_landau_n_pair_th4))). Time Defined. (* constant 1044 *) Definition l_e_st_eq_landau_n_second1is2 : (forall (alpha:Type), (forall (a:alpha), (forall (b:alpha), l_e_is alpha b (l_e_st_eq_landau_n_second1 alpha (l_e_st_eq_landau_n_pair1 alpha a b))))). exact (fun (alpha:Type) => (fun (a:alpha) => (fun (b:alpha) => l_e_symis alpha (l_e_st_eq_landau_n_second1 alpha (l_e_st_eq_landau_n_pair1 alpha a b)) b (l_e_st_eq_landau_n_second1is1 alpha a b)))). Time Defined. (* constant 1045 *) Definition l_e_st_eq_landau_n_pair_t11 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), (forall (q:l_e_st_eq_landau_n_pair1type alpha), (forall (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)), (forall (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (u1:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t), l_e_is alpha (p u) (l_e_st_eq_landau_n_first1 alpha p)))))))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => (fun (q:l_e_st_eq_landau_n_pair1type alpha) => (fun (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)) => (fun (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (u1:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t) => l_e_isf (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) alpha p u l_e_st_eq_landau_n_1t u1))))))). Time Defined. (* constant 1046 *) Definition l_e_st_eq_landau_n_pair_t12 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), (forall (q:l_e_st_eq_landau_n_pair1type alpha), (forall (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)), (forall (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (u1:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t), l_e_is alpha (l_e_st_eq_landau_n_first1 alpha q) (q u)))))))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => (fun (q:l_e_st_eq_landau_n_pair1type alpha) => (fun (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)) => (fun (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (u1:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t) => l_e_symis alpha (q u) (l_e_st_eq_landau_n_first1 alpha q) (l_e_isf (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) alpha q u l_e_st_eq_landau_n_1t u1)))))))). Time Defined. (* constant 1047 *) Definition l_e_st_eq_landau_n_pair_t13 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), (forall (q:l_e_st_eq_landau_n_pair1type alpha), (forall (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)), (forall (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (u1:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t), l_e_is alpha (p u) (q u)))))))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => (fun (q:l_e_st_eq_landau_n_pair1type alpha) => (fun (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)) => (fun (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (u1:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t) => l_e_tr3is alpha (p u) (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q) (q u) (l_e_st_eq_landau_n_pair_t11 alpha p q i j u u1) i (l_e_st_eq_landau_n_pair_t12 alpha p q i j u u1)))))))). Time Defined. (* constant 1048 *) Definition l_e_st_eq_landau_n_pair_t14 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), (forall (q:l_e_st_eq_landau_n_pair1type alpha), (forall (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)), (forall (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (u2:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t), l_e_is alpha (p u) (l_e_st_eq_landau_n_second1 alpha p)))))))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => (fun (q:l_e_st_eq_landau_n_pair1type alpha) => (fun (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)) => (fun (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (u2:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t) => l_e_isf (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) alpha p u l_e_st_eq_landau_n_2t u2))))))). Time Defined. (* constant 1049 *) Definition l_e_st_eq_landau_n_pair_t15 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), (forall (q:l_e_st_eq_landau_n_pair1type alpha), (forall (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)), (forall (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (u2:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t), l_e_is alpha (l_e_st_eq_landau_n_second1 alpha q) (q u)))))))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => (fun (q:l_e_st_eq_landau_n_pair1type alpha) => (fun (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)) => (fun (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (u2:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t) => l_e_symis alpha (q u) (l_e_st_eq_landau_n_second1 alpha q) (l_e_isf (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) alpha q u l_e_st_eq_landau_n_2t u2)))))))). Time Defined. (* constant 1050 *) Definition l_e_st_eq_landau_n_pair_t16 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), (forall (q:l_e_st_eq_landau_n_pair1type alpha), (forall (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)), (forall (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), (forall (u2:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t), l_e_is alpha (p u) (q u)))))))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => (fun (q:l_e_st_eq_landau_n_pair1type alpha) => (fun (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)) => (fun (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => (fun (u2:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t) => l_e_tr3is alpha (p u) (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q) (q u) (l_e_st_eq_landau_n_pair_t14 alpha p q i j u u2) j (l_e_st_eq_landau_n_pair_t15 alpha p q i j u u2)))))))). Time Defined. (* constant 1051 *) Definition l_e_st_eq_landau_n_pair_t17 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), (forall (q:l_e_st_eq_landau_n_pair1type alpha), (forall (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)), (forall (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_is alpha (p u) (q u))))))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => (fun (q:l_e_st_eq_landau_n_pair1type alpha) => (fun (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)) => (fun (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => l_orapp (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t) (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t) (l_e_is alpha (p u) (q u)) (l_e_st_eq_landau_n_pair_th3 u) (fun (t:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_1t) => l_e_st_eq_landau_n_pair_t13 alpha p q i j u t) (fun (t:l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) u l_e_st_eq_landau_n_2t) => l_e_st_eq_landau_n_pair_t16 alpha p q i j u t))))))). Time Defined. (* constant 1052 *) Definition l_e_st_eq_landau_n_pair_th5 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), (forall (q:l_e_st_eq_landau_n_pair1type alpha), (forall (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)), (forall (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)), l_e_is (l_e_st_eq_landau_n_pair1type alpha) p q))))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => (fun (q:l_e_st_eq_landau_n_pair1type alpha) => (fun (i:l_e_is alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_first1 alpha q)) => (fun (j:l_e_is alpha (l_e_st_eq_landau_n_second1 alpha p) (l_e_st_eq_landau_n_second1 alpha q)) => l_e_fisi (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) alpha p q (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) => l_e_st_eq_landau_n_pair_t17 alpha p q i j t)))))). Time Defined. (* constant 1053 *) Definition l_e_st_eq_landau_n_pair_q0 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), l_e_st_eq_landau_n_pair1type alpha)). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => l_e_st_eq_landau_n_pair1 alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_second1 alpha p))). Time Defined. (* constant 1054 *) Definition l_e_st_eq_landau_n_pair_t18 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), l_e_is alpha (l_e_st_eq_landau_n_first1 alpha (l_e_st_eq_landau_n_pair_q0 alpha p)) (l_e_st_eq_landau_n_first1 alpha p))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => l_e_st_eq_landau_n_first1is1 alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_second1 alpha p))). Time Defined. (* constant 1055 *) Definition l_e_st_eq_landau_n_pair_t19 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), l_e_is alpha (l_e_st_eq_landau_n_second1 alpha (l_e_st_eq_landau_n_pair_q0 alpha p)) (l_e_st_eq_landau_n_second1 alpha p))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => l_e_st_eq_landau_n_second1is1 alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_second1 alpha p))). Time Defined. (* constant 1056 *) Definition l_e_st_eq_landau_n_pair1is1 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), l_e_is (l_e_st_eq_landau_n_pair1type alpha) (l_e_st_eq_landau_n_pair1 alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_second1 alpha p)) p)). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => l_e_st_eq_landau_n_pair_th5 alpha (l_e_st_eq_landau_n_pair_q0 alpha p) p (l_e_st_eq_landau_n_pair_t18 alpha p) (l_e_st_eq_landau_n_pair_t19 alpha p))). Time Defined. (* constant 1057 *) Definition l_e_st_eq_landau_n_pair1is2 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), l_e_is (l_e_st_eq_landau_n_pair1type alpha) p (l_e_st_eq_landau_n_pair1 alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_second1 alpha p)))). exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => l_e_symis (l_e_st_eq_landau_n_pair1type alpha) (l_e_st_eq_landau_n_pair1 alpha (l_e_st_eq_landau_n_first1 alpha p) (l_e_st_eq_landau_n_second1 alpha p)) p (l_e_st_eq_landau_n_pair1is1 alpha p))). Time Defined. (* constant 1058 *) Definition l_e_st_eq_landau_n_lessisi3 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lessis x x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessisi2 x x (l_e_refis l_e_st_eq_landau_n_nat x)). Time Defined. (* constant 1059 *) Definition l_e_st_eq_landau_n_1out : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_1to x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_outn x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a x)). Time Defined. (* constant 1060 *) Definition l_e_st_eq_landau_n_xout : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_1to x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_outn x x (l_e_st_eq_landau_n_lessisi3 x)). Time Defined. (* constant 1061 *) Definition l_e_st_eq_landau_n_left_ui : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_nat)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_inn y u)))). Time Defined. (* constant 1062 *) Definition l_e_st_eq_landau_n_left_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_left_ui x y l u) y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_1top y u)))). Time Defined. (* constant 1063 *) Definition l_e_st_eq_landau_n_left_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_left_ui x y l u) x)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_left_ui x y l u) y x (l_e_st_eq_landau_n_left_t1 x y l u) l)))). Time Defined. (* constant 1064 *) Definition l_e_st_eq_landau_n_left1to : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_1to x)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_left_ui x y l u) (l_e_st_eq_landau_n_left_t2 x y l u))))). Time Defined. (* constant 1065 *) Definition l_e_st_eq_landau_n_left_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (u:l_e_st_eq_landau_n_1to y), (forall (v:l_e_st_eq_landau_n_1to y), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y l u) (l_e_st_eq_landau_n_left1to x y l v)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_left_ui x y l u) (l_e_st_eq_landau_n_left_ui x y l v))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (u:l_e_st_eq_landau_n_1to y) => (fun (v:l_e_st_eq_landau_n_1to y) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y l u) (l_e_st_eq_landau_n_left1to x y l v)) => l_e_st_eq_landau_n_isoutne x (l_e_st_eq_landau_n_left_ui x y l u) (l_e_st_eq_landau_n_left_t2 x y l u) (l_e_st_eq_landau_n_left_ui x y l v) (l_e_st_eq_landau_n_left_t2 x y l v) i)))))). Time Defined. (* constant 1066 *) Definition l_e_st_eq_landau_n_thleft1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (u:l_e_st_eq_landau_n_1to y), (forall (v:l_e_st_eq_landau_n_1to y), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y l u) (l_e_st_eq_landau_n_left1to x y l v)), l_e_is (l_e_st_eq_landau_n_1to y) u v)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (u:l_e_st_eq_landau_n_1to y) => (fun (v:l_e_st_eq_landau_n_1to y) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y l u) (l_e_st_eq_landau_n_left1to x y l v)) => l_e_st_eq_landau_n_isinne y u v (l_e_st_eq_landau_n_left_t3 x y l u v i))))))). Time Defined. (* constant 1067 *) Definition l_e_st_eq_landau_n_thleft2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), l_e_injective (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_left1to x y l t)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (u:l_e_st_eq_landau_n_1to y) => (fun (v:l_e_st_eq_landau_n_1to y) => (fun (t:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y l u) (l_e_st_eq_landau_n_left1to x y l v)) => l_e_st_eq_landau_n_thleft1 x y l u v t)))))). Time Defined. (* constant 1068 *) Definition l_e_st_eq_landau_n_right_ui : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_nat))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_inn y u))). Time Defined. (* constant 1069 *) Definition l_e_st_eq_landau_n_right_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_right_ui x y u) y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_1top y u))). Time Defined. (* constant 1070 *) Definition l_e_st_eq_landau_n_right_t5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_right_ui x y u)) (l_e_st_eq_landau_n_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_right_ui x y u) y x (l_e_st_eq_landau_n_right_t4 x y u)))). Time Defined. (* constant 1071 *) Definition l_e_st_eq_landau_n_right1to : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_right_ui x y u)) (l_e_st_eq_landau_n_right_t5 x y u)))). Time Defined. (* constant 1072 *) Definition l_e_st_eq_landau_n_right_t6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_1to y), (forall (v:l_e_st_eq_landau_n_1to y), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_right1to x y u) (l_e_st_eq_landau_n_right1to x y v)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_right_ui x y u)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_right_ui x y v))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_1to y) => (fun (v:l_e_st_eq_landau_n_1to y) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_right1to x y u) (l_e_st_eq_landau_n_right1to x y v)) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_right_ui x y u)) (l_e_st_eq_landau_n_right_t5 x y u) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_right_ui x y v)) (l_e_st_eq_landau_n_right_t5 x y v) i))))). Time Defined. (* constant 1073 *) Definition l_e_st_eq_landau_n_right_t7 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_1to y), (forall (v:l_e_st_eq_landau_n_1to y), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_right1to x y u) (l_e_st_eq_landau_n_right1to x y v)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_right_ui x y u) (l_e_st_eq_landau_n_right_ui x y v)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_1to y) => (fun (v:l_e_st_eq_landau_n_1to y) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_right1to x y u) (l_e_st_eq_landau_n_right1to x y v)) => l_e_st_eq_landau_n_satz20e (l_e_st_eq_landau_n_right_ui x y u) (l_e_st_eq_landau_n_right_ui x y v) x (l_e_st_eq_landau_n_right_t6 x y u v i)))))). Time Defined. (* constant 1074 *) Definition l_e_st_eq_landau_n_thright1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_1to y), (forall (v:l_e_st_eq_landau_n_1to y), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_right1to x y u) (l_e_st_eq_landau_n_right1to x y v)), l_e_is (l_e_st_eq_landau_n_1to y) u v))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_1to y) => (fun (v:l_e_st_eq_landau_n_1to y) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_right1to x y u) (l_e_st_eq_landau_n_right1to x y v)) => l_e_st_eq_landau_n_isinne y u v (l_e_st_eq_landau_n_right_t7 x y u v i)))))). Time Defined. (* constant 1075 *) Definition l_e_st_eq_landau_n_left : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), alpha)), (forall (t:l_e_st_eq_landau_n_1to y), alpha)))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), alpha)) => (fun (t:l_e_st_eq_landau_n_1to y) => f (l_e_st_eq_landau_n_left1to x y l t))))))). Time Defined. (* constant 1076 *) Definition l_e_st_eq_landau_n_right : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), alpha)), (forall (t:l_e_st_eq_landau_n_1to y), alpha))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), alpha)) => (fun (t:l_e_st_eq_landau_n_1to y) => f (l_e_st_eq_landau_n_right1to x y t)))))). Time Defined. (* constant 1077 *) Definition l_e_st_eq_landau_n_left_t4 : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)), l_e_st_eq_landau_n_lessis y x))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)) => l_e_st_eq_landau_n_lessisi2 y x i))))). Time Defined. (* constant 1078 *) Definition l_e_st_eq_landau_n_left_t5 : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)), l_e_st_eq_landau_n_lessis x y))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)) => l_e_st_eq_landau_n_lessisi2 x y (l_e_symis l_e_st_eq_landau_n_nat y x i)))))). Time Defined. (* constant 1079 *) Definition l_e_st_eq_landau_n_left_f1 : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)), (forall (t:l_e_st_eq_landau_n_1to x), alpha)))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)) => l_e_st_eq_landau_n_left alpha y x (l_e_st_eq_landau_n_left_t5 alpha x y i f) f))))). Time Defined. (* constant 1080 *) Definition l_e_st_eq_landau_n_left_f2 : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)), (forall (t:l_e_st_eq_landau_n_1to y), alpha)))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)) => l_e_st_eq_landau_n_left alpha x y (l_e_st_eq_landau_n_left_t4 alpha x y i f) (l_e_st_eq_landau_n_left_f1 alpha x y i f)))))). Time Defined. (* constant 1081 *) Definition l_e_st_eq_landau_n_left_t6 : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)), (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_inn y u) (l_e_st_eq_landau_n_inn x (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_left_t4 alpha x y i f) u)))))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_inn y u) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_inn y u) y x (l_e_st_eq_landau_n_1top y u) (l_e_st_eq_landau_n_left_t4 alpha x y i f)))))))). Time Defined. (* constant 1082 *) Definition l_e_st_eq_landau_n_left_t7 : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)), (forall (u:l_e_st_eq_landau_n_1to y), l_e_is (l_e_st_eq_landau_n_1to y) u (l_e_st_eq_landau_n_left1to y x (l_e_st_eq_landau_n_left_t5 alpha x y i f) (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_left_t4 alpha x y i f) u)))))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_tris (l_e_st_eq_landau_n_1to y) u (l_e_st_eq_landau_n_outn y (l_e_st_eq_landau_n_inn y u) (l_e_st_eq_landau_n_1top y u)) (l_e_st_eq_landau_n_left1to y x (l_e_st_eq_landau_n_left_t5 alpha x y i f) (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_left_t4 alpha x y i f) u)) (l_e_st_eq_landau_n_isoutinn y u) (l_e_st_eq_landau_n_isoutni y (l_e_st_eq_landau_n_inn y u) (l_e_st_eq_landau_n_1top y u) (l_e_st_eq_landau_n_inn x (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_left_t4 alpha x y i f) u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_inn x (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_left_t4 alpha x y i f) u)) x y (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_left_t4 alpha x y i f) u)) (l_e_st_eq_landau_n_left_t5 alpha x y i f)) (l_e_st_eq_landau_n_left_t6 alpha x y i f u)))))))). Time Defined. (* constant 1083 *) Definition l_e_st_eq_landau_n_left_t8 : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)), (forall (u:l_e_st_eq_landau_n_1to y), l_e_is alpha (f u) (l_e_st_eq_landau_n_left_f2 alpha x y i f u))))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)) => (fun (u:l_e_st_eq_landau_n_1to y) => l_e_isf (l_e_st_eq_landau_n_1to y) alpha f u (l_e_st_eq_landau_n_left1to y x (l_e_st_eq_landau_n_left_t5 alpha x y i f) (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_left_t4 alpha x y i f) u)) (l_e_st_eq_landau_n_left_t7 alpha x y i f u))))))). Time Defined. (* constant 1084 *) Definition l_e_st_eq_landau_n_thleft : (forall (alpha:Type), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)), l_e_is (forall (t:l_e_st_eq_landau_n_1to y), alpha) f (l_e_st_eq_landau_n_left alpha x y (l_e_st_eq_landau_n_lessisi2 y x i) (l_e_st_eq_landau_n_left alpha y x (l_e_st_eq_landau_n_lessisi2 x y (l_e_symis l_e_st_eq_landau_n_nat y x i)) f))))))). exact (fun (alpha:Type) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to y), alpha)) => l_e_fisi (l_e_st_eq_landau_n_1to y) alpha f (l_e_st_eq_landau_n_left_f2 alpha x y i f) (fun (t:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_left_t8 alpha x y i f t)))))). Time Defined. (* constant 1085 *) Definition l_e_st_eq_landau_n_frac : Type. exact (l_e_st_eq_landau_n_pair1type l_e_st_eq_landau_n_nat). Time Defined. (* constant 1086 *) Definition l_e_st_eq_landau_n_fr : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_frac)). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pair1 l_e_st_eq_landau_n_nat x1 x2)). Time Defined. (* constant 1087 *) Definition l_e_st_eq_landau_n_num : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_first1 l_e_st_eq_landau_n_nat x). Time Defined. (* constant 1088 *) Definition l_e_st_eq_landau_n_den : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_second1 l_e_st_eq_landau_n_nat x). Time Defined. (* constant 1089 *) Definition l_e_st_eq_landau_n_numis : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) x1)). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_first1is1 l_e_st_eq_landau_n_nat x1 x2)). Time Defined. (* constant 1090 *) Definition l_e_st_eq_landau_n_isnum : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x1 (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_first1is2 l_e_st_eq_landau_n_nat x1 x2)). Time Defined. (* constant 1091 *) Definition l_e_st_eq_landau_n_denis : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)) x2)). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_second1is1 l_e_st_eq_landau_n_nat x1 x2)). Time Defined. (* constant 1092 *) Definition l_e_st_eq_landau_n_isden : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x2 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_second1is2 l_e_st_eq_landau_n_nat x1 x2)). Time Defined. (* constant 1093 *) Definition l_e_st_eq_landau_n_1x : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_num x). Time Defined. (* constant 1094 *) Definition l_e_st_eq_landau_n_2x : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_den x). Time Defined. (* constant 1095 *) Definition l_e_st_eq_landau_n_fris : (forall (x:l_e_st_eq_landau_n_frac), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) x). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_pair1is1 l_e_st_eq_landau_n_nat x). Time Defined. (* constant 1096 *) Definition l_e_st_eq_landau_n_isfr : (forall (x:l_e_st_eq_landau_n_frac), l_e_is l_e_st_eq_landau_n_frac x (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x))). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_pair1is2 l_e_st_eq_landau_n_nat x). Time Defined. (* constant 1097 *) Definition l_e_st_eq_landau_n_12isnd : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists12 x1 (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) y2 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_isnum x1 x2) (l_e_st_eq_landau_n_isden y1 y2))))). Time Defined. (* constant 1098 *) Definition l_e_st_eq_landau_n_ndis12 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts x1 y2))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_12isnd x1 x2 y1 y2))))). Time Defined. (* constant 1099 *) Definition l_e_st_eq_landau_n_1disnd : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists1 n1 (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_isnum n1 n2)))). Time Defined. (* constant 1100 *) Definition l_e_st_eq_landau_n_ndis1d : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_1disnd x n1 n2)))). Time Defined. (* constant 1101 *) Definition l_e_st_eq_landau_n_n2isnd : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists2 n2 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_isden n1 n2)))). Time Defined. (* constant 1102 *) Definition l_e_st_eq_landau_n_ndisn2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_n2isnd x n1 n2)))). Time Defined. (* constant 1103 *) Definition l_e_st_eq_landau_n_isn : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x1 n), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr n x2))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x1 n) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_fr t x2) x1 n i)))). Time Defined. (* constant 1104 *) Definition l_e_st_eq_landau_n_isd : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x2 n), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr x1 n))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x2 n) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_fr x1 t) x2 n i)))). Time Defined. (* constant 1105 *) Definition l_e_st_eq_landau_n_isnd : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x1 y1), (forall (j:l_e_st_eq_landau_n_is x2 y2), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2))))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x1 y1) => (fun (j:l_e_st_eq_landau_n_is x2 y2) => l_e_tris l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 x2) (l_e_st_eq_landau_n_fr y1 y2) (l_e_st_eq_landau_n_isn x1 x2 y1 i) (l_e_st_eq_landau_n_isd y1 x2 y2 j))))))). Time Defined. (* constant 1106 *) Definition l_e_st_eq_landau_n_1y : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_num y)). Time Defined. (* constant 1107 *) Definition l_e_st_eq_landau_n_2y : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_den y)). Time Defined. (* constant 1108 *) Definition l_e_st_eq_landau_n_eq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), Prop)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1109 *) Definition l_e_st_eq_landau_n_eqi12 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_ndis12 x1 x2 y1 y2) i (l_e_st_eq_landau_n_12isnd y1 y2 x1 x2)))))). Time Defined. (* constant 1110 *) Definition l_e_st_eq_landau_n_eqi1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))), l_e_st_eq_landau_n_eq x (l_e_st_eq_landau_n_fr n1 n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) => l_e_isp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq t (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) x (l_e_st_eq_landau_n_eqi12 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x) n1 n2 i) (l_e_st_eq_landau_n_fris x))))). Time Defined. (* constant 1111 *) Definition l_e_st_eq_landau_n_eqi2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr n1 n2) x)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) => l_e_isp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr n1 n2) t) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) x (l_e_st_eq_landau_n_eqi12 n1 n2 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x) i) (l_e_st_eq_landau_n_fris x))))). Time Defined. (* constant 1112 *) Definition l_e_st_eq_landau_n_satz37 : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq x x). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_refis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x))). Time Defined. (* constant 1113 *) Definition l_e_st_eq_landau_n_refeq : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq x x). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz37 x). Time Defined. (* constant 1114 *) Definition l_e_st_eq_landau_n_refeq1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (i:l_e_is l_e_st_eq_landau_n_frac x y), l_e_st_eq_landau_n_eq x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (i:l_e_is l_e_st_eq_landau_n_frac x y) => l_e_isp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq x t) x y (l_e_st_eq_landau_n_refeq x) i))). Time Defined. (* constant 1115 *) Definition l_e_st_eq_landau_n_refeq2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (i:l_e_is l_e_st_eq_landau_n_frac x y), l_e_st_eq_landau_n_eq y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (i:l_e_is l_e_st_eq_landau_n_frac x y) => l_e_isp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq t x) x y (l_e_st_eq_landau_n_refeq x) i))). Time Defined. (* constant 1116 *) Definition l_e_st_eq_landau_n_eqnd : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x1 y1), (forall (j:l_e_st_eq_landau_n_is x2 y2), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2))))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x1 y1) => (fun (j:l_e_st_eq_landau_n_is x2 y2) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2) (l_e_st_eq_landau_n_isnd x1 x2 y1 y2 i j))))))). Time Defined. (* constant 1117 *) Definition l_e_st_eq_landau_n_eqn : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x1 n), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr n x2))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x1 n) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr n x2) (l_e_st_eq_landau_n_isn x1 x2 n i))))). Time Defined. (* constant 1118 *) Definition l_e_st_eq_landau_n_eqd : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x2 n), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr x1 n))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x2 n) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr x1 n) (l_e_st_eq_landau_n_isd x1 x2 n i))))). Time Defined. (* constant 1119 *) Definition l_e_st_eq_landau_n_satz38 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) e))). Time Defined. (* constant 1120 *) Definition l_e_st_eq_landau_n_symeq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz38 x y e))). Time Defined. (* constant 1121 *) Definition l_e_st_eq_landau_n_ii1_t1 : (forall (a:l_e_st_eq_landau_n_nat), (forall (b:l_e_st_eq_landau_n_nat), (forall (c:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts d (l_e_st_eq_landau_n_ts b c)))))). exact (fun (a:l_e_st_eq_landau_n_nat) => (fun (b:l_e_st_eq_landau_n_nat) => (fun (c:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts b c) d) (l_e_st_eq_landau_n_ts d (l_e_st_eq_landau_n_ts b c)) (l_e_st_eq_landau_n_assts2 b c d) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_ts b c) d))))). Time Defined. (* constant 1122 *) Definition l_e_st_eq_landau_n_stets : (forall (a:l_e_st_eq_landau_n_nat), (forall (b:l_e_st_eq_landau_n_nat), (forall (c:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a d) (l_e_st_eq_landau_n_ts c b)))))). exact (fun (a:l_e_st_eq_landau_n_nat) => (fun (b:l_e_st_eq_landau_n_nat) => (fun (c:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_nat) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts a (l_e_st_eq_landau_n_ts b (l_e_st_eq_landau_n_ts c d))) (l_e_st_eq_landau_n_ts a (l_e_st_eq_landau_n_ts d (l_e_st_eq_landau_n_ts b c))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a d) (l_e_st_eq_landau_n_ts b c)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a d) (l_e_st_eq_landau_n_ts c b)) (l_e_st_eq_landau_n_assts1 a b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts d (l_e_st_eq_landau_n_ts b c)) a (l_e_st_eq_landau_n_ii1_t1 a b c d)) (l_e_st_eq_landau_n_assts2 a d (l_e_st_eq_landau_n_ts b c)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts b c) (l_e_st_eq_landau_n_ts c b) (l_e_st_eq_landau_n_ts a d) (l_e_st_eq_landau_n_comts b c)))))). Time Defined. (* constant 1123 *) Definition l_e_st_eq_landau_n_ii1_t2 : (forall (a:l_e_st_eq_landau_n_nat), (forall (b:l_e_st_eq_landau_n_nat), (forall (c:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts d (l_e_st_eq_landau_n_ts c b)))))). exact (fun (a:l_e_st_eq_landau_n_nat) => (fun (b:l_e_st_eq_landau_n_nat) => (fun (c:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts c d) b) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts d c) b) (l_e_st_eq_landau_n_ts d (l_e_st_eq_landau_n_ts c b)) (l_e_st_eq_landau_n_comts b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ists1 (l_e_st_eq_landau_n_ts c d) (l_e_st_eq_landau_n_ts d c) b (l_e_st_eq_landau_n_comts c d)) (l_e_st_eq_landau_n_assts1 d c b))))). Time Defined. (* constant 1124 *) Definition l_e_st_eq_landau_n_ii1_anders : (forall (a:l_e_st_eq_landau_n_nat), (forall (b:l_e_st_eq_landau_n_nat), (forall (c:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a d) (l_e_st_eq_landau_n_ts c b)))))). exact (fun (a:l_e_st_eq_landau_n_nat) => (fun (b:l_e_st_eq_landau_n_nat) => (fun (c:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts a (l_e_st_eq_landau_n_ts b (l_e_st_eq_landau_n_ts c d))) (l_e_st_eq_landau_n_ts a (l_e_st_eq_landau_n_ts d (l_e_st_eq_landau_n_ts c b))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a d) (l_e_st_eq_landau_n_ts c b)) (l_e_st_eq_landau_n_assts1 a b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts b (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts d (l_e_st_eq_landau_n_ts c b)) a (l_e_st_eq_landau_n_ii1_t2 a b c d)) (l_e_st_eq_landau_n_assts2 a d (l_e_st_eq_landau_n_ts c b)))))). Time Defined. (* constant 1125 *) Definition l_e_st_eq_landau_n_1z : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_num z))). Time Defined. (* constant 1126 *) Definition l_e_st_eq_landau_n_2z : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_den z))). Time Defined. (* constant 1127 *) Definition l_e_st_eq_landau_n_139_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_e_st_eq_landau_n_ists12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) e f))))). Time Defined. (* constant 1128 *) Definition l_e_st_eq_landau_n_139_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)))))). Time Defined. (* constant 1129 *) Definition l_e_st_eq_landau_n_139_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)))))))). Time Defined. (* constant 1130 *) Definition l_e_st_eq_landau_n_139_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_139_t2 x y z e f)) (l_e_st_eq_landau_n_139_t1 x y z e f) (l_e_st_eq_landau_n_139_t3 x y z e f)))))). Time Defined. (* constant 1131 *) Definition l_e_st_eq_landau_n_satz39 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_eq x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_e_st_eq_landau_n_satz33b (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_139_t4 x y z e f)))))). Time Defined. (* constant 1132 *) Definition l_e_st_eq_landau_n_139_anders : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ists12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) e f) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)))))))). Time Defined. (* constant 1133 *) Definition l_e_st_eq_landau_n_treq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_eq x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_e_st_eq_landau_n_satz39 x y z e f))))). Time Defined. (* constant 1134 *) Definition l_e_st_eq_landau_n_treq1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq z x), (forall (f:l_e_st_eq_landau_n_eq z y), l_e_st_eq_landau_n_eq x y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq z x) => (fun (f:l_e_st_eq_landau_n_eq z y) => l_e_st_eq_landau_n_treq x z y (l_e_st_eq_landau_n_symeq z x e) f))))). Time Defined. (* constant 1135 *) Definition l_e_st_eq_landau_n_treq2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_eq x y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_e_st_eq_landau_n_treq x z y e (l_e_st_eq_landau_n_symeq y z f)))))). Time Defined. (* constant 1136 *) Definition l_e_st_eq_landau_n_tr3eq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), (forall (g:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_eq x u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => (fun (g:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_treq x y u e (l_e_st_eq_landau_n_treq y z u f g)))))))). Time Defined. (* constant 1137 *) Definition l_e_st_eq_landau_n_tr4eq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (v:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), (forall (g:l_e_st_eq_landau_n_eq z u), (forall (h:l_e_st_eq_landau_n_eq u v), l_e_st_eq_landau_n_eq x v))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => (fun (g:l_e_st_eq_landau_n_eq z u) => (fun (h:l_e_st_eq_landau_n_eq u v) => l_e_st_eq_landau_n_tr3eq x y z v e f (l_e_st_eq_landau_n_treq z u v g h)))))))))). Time Defined. (* constant 1138 *) Definition l_e_st_eq_landau_n_satz40 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq x (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_eqi1 x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n) (l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts n (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n) (l_e_st_eq_landau_n_ts n (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2x x) n)) (l_e_st_eq_landau_n_assts2 (l_e_st_eq_landau_n_1x x) n (l_e_st_eq_landau_n_2x x))))). Time Defined. (* constant 1139 *) Definition l_e_st_eq_landau_n_satz40a : (forall (x:l_e_st_eq_landau_n_frac), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n)) x)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_symeq x (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n)) (l_e_st_eq_landau_n_satz40 x n))). Time Defined. (* constant 1140 *) Definition l_e_st_eq_landau_n_satz40b : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x1 n) (l_e_st_eq_landau_n_ts x2 n))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_eqi12 x1 x2 (l_e_st_eq_landau_n_ts x1 n) (l_e_st_eq_landau_n_ts x2 n) (l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x1 (l_e_st_eq_landau_n_ts x2 n)) (l_e_st_eq_landau_n_ts x1 (l_e_st_eq_landau_n_ts n x2)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x1 n) x2) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts x2 n) (l_e_st_eq_landau_n_ts n x2) x1 (l_e_st_eq_landau_n_comts x2 n)) (l_e_st_eq_landau_n_assts2 x1 n x2))))). Time Defined. (* constant 1141 *) Definition l_e_st_eq_landau_n_satz40c : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x1 n) (l_e_st_eq_landau_n_ts x2 n)) (l_e_st_eq_landau_n_fr x1 x2)))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_symeq (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x1 n) (l_e_st_eq_landau_n_ts x2 n)) (l_e_st_eq_landau_n_satz40b x1 x2 n)))). Time Defined. (* constant 1142 *) Definition l_e_st_eq_landau_n_moref : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), Prop)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1143 *) Definition l_e_st_eq_landau_n_lessf : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), Prop)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1144 *) Definition l_e_st_eq_landau_n_morefi12 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts y1 x2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_12isnd x1 x2 y1 y2) (l_e_st_eq_landau_n_12isnd y1 y2 x1 x2) m))))). Time Defined. (* constant 1145 *) Definition l_e_st_eq_landau_n_lessfi12 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts y1 x2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_12isnd x1 x2 y1 y2) (l_e_st_eq_landau_n_12isnd y1 y2 x1 x2) l))))). Time Defined. (* constant 1146 *) Definition l_e_st_eq_landau_n_morefi1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))), l_e_st_eq_landau_n_moref x (l_e_st_eq_landau_n_fr n1 n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_n2isnd x n1 n2) (l_e_st_eq_landau_n_1disnd x n1 n2) m)))). Time Defined. (* constant 1147 *) Definition l_e_st_eq_landau_n_morefi2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr n1 n2) x)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_1disnd x n1 n2) (l_e_st_eq_landau_n_n2isnd x n1 n2) m)))). Time Defined. (* constant 1148 *) Definition l_e_st_eq_landau_n_lessfi1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))), l_e_st_eq_landau_n_lessf x (l_e_st_eq_landau_n_fr n1 n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_n2isnd x n1 n2) (l_e_st_eq_landau_n_1disnd x n1 n2) l)))). Time Defined. (* constant 1149 *) Definition l_e_st_eq_landau_n_lessfi2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr n1 n2) x)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_1disnd x n1 n2) (l_e_st_eq_landau_n_n2isnd x n1 n2) l)))). Time Defined. (* constant 1150 *) Definition l_e_st_eq_landau_n_satz41 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_orec3 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz10 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1151 *) Definition l_e_st_eq_landau_n_satz41a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_or3 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz10a (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1152 *) Definition l_e_st_eq_landau_n_satz41b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_ec3 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1153 *) Definition l_e_st_eq_landau_n_satz42 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_lessf y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz11 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) m))). Time Defined. (* constant 1154 *) Definition l_e_st_eq_landau_n_satz43 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_moref y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) l))). Time Defined. (* constant 1155 *) Definition l_e_st_eq_landau_n_1u : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_num u)))). Time Defined. (* constant 1156 *) Definition l_e_st_eq_landau_n_2u : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_den u)))). Time Defined. (* constant 1157 *) Definition l_e_st_eq_landau_n_244_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => l_e_st_eq_landau_n_ists12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) f (l_e_st_eq_landau_n_symeq x z e)))))))). Time Defined. (* constant 1158 *) Definition l_e_st_eq_landau_n_244_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_244_t1 x y z u m e f) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z))))))))). Time Defined. (* constant 1159 *) Definition l_e_st_eq_landau_n_244_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_244_t2 x y z u m e f)) (l_e_st_eq_landau_n_satz32d (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) m)))))))). Time Defined. (* constant 1160 *) Definition l_e_st_eq_landau_n_satz44 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), l_e_st_eq_landau_n_moref z u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => l_e_st_eq_landau_n_satz33a (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_244_t3 x y z u m e f))))))))). Time Defined. (* constant 1161 *) Definition l_e_st_eq_landau_n_eqmoref12 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), (forall (m:l_e_st_eq_landau_n_moref x z), l_e_st_eq_landau_n_moref y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => (fun (m:l_e_st_eq_landau_n_moref x z) => l_e_st_eq_landau_n_satz44 x z y u m e f))))))). Time Defined. (* constant 1162 *) Definition l_e_st_eq_landau_n_eqmoref1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (m:l_e_st_eq_landau_n_moref x z), l_e_st_eq_landau_n_moref y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (m:l_e_st_eq_landau_n_moref x z) => l_e_st_eq_landau_n_satz44 x z y z m e (l_e_st_eq_landau_n_refeq z)))))). Time Defined. (* constant 1163 *) Definition l_e_st_eq_landau_n_eqmoref2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (m:l_e_st_eq_landau_n_moref z x), l_e_st_eq_landau_n_moref z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (m:l_e_st_eq_landau_n_moref z x) => l_e_st_eq_landau_n_satz44 z x z y m (l_e_st_eq_landau_n_refeq z) e))))). Time Defined. (* constant 1164 *) Definition l_e_st_eq_landau_n_satz45 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), l_e_st_eq_landau_n_lessf z u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => l_e_st_eq_landau_n_satz42 u z (l_e_st_eq_landau_n_satz44 y x u z (l_e_st_eq_landau_n_satz43 x y l) f e)))))))). Time Defined. (* constant 1165 *) Definition l_e_st_eq_landau_n_eqlessf12 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), (forall (l:l_e_st_eq_landau_n_lessf x z), l_e_st_eq_landau_n_lessf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => (fun (l:l_e_st_eq_landau_n_lessf x z) => l_e_st_eq_landau_n_satz45 x z y u l e f))))))). Time Defined. (* constant 1166 *) Definition l_e_st_eq_landau_n_eqlessf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (l:l_e_st_eq_landau_n_lessf x z), l_e_st_eq_landau_n_lessf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (l:l_e_st_eq_landau_n_lessf x z) => l_e_st_eq_landau_n_satz45 x z y z l e (l_e_st_eq_landau_n_refeq z)))))). Time Defined. (* constant 1167 *) Definition l_e_st_eq_landau_n_eqlessf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (l:l_e_st_eq_landau_n_lessf z x), l_e_st_eq_landau_n_lessf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (l:l_e_st_eq_landau_n_lessf z x) => l_e_st_eq_landau_n_satz45 z x z y l (l_e_st_eq_landau_n_refeq z) e))))). Time Defined. (* constant 1168 *) Definition l_e_st_eq_landau_n_moreq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), Prop)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_or (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y))). Time Defined. (* constant 1169 *) Definition l_e_st_eq_landau_n_lesseq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), Prop)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_or (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y))). Time Defined. (* constant 1170 *) Definition l_e_st_eq_landau_n_moreqi2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_moreq x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_ori2 (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) e))). Time Defined. (* constant 1171 *) Definition l_e_st_eq_landau_n_lesseqi2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_lesseq x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_ori2 (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) e))). Time Defined. (* constant 1172 *) Definition l_e_st_eq_landau_n_moreqi1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moreq x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_ori1 (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) m))). Time Defined. (* constant 1173 *) Definition l_e_st_eq_landau_n_lesseqi1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lesseq x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_ori1 (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) l))). Time Defined. (* constant 1174 *) Definition l_e_st_eq_landau_n_satz41c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), l_not (l_e_st_eq_landau_n_lessf x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => l_ec3_th7 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41b x y) (l_comor (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) m)))). Time Defined. (* constant 1175 *) Definition l_e_st_eq_landau_n_satz41d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), l_not (l_e_st_eq_landau_n_moref x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => l_ec3_th9 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41b x y) l))). Time Defined. (* constant 1176 *) Definition l_e_st_eq_landau_n_satz41e : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (n:l_not (l_e_st_eq_landau_n_moref x y)), l_e_st_eq_landau_n_lesseq x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (n:l_not (l_e_st_eq_landau_n_moref x y)) => l_or3_th2 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41a x y) n))). Time Defined. (* constant 1177 *) Definition l_e_st_eq_landau_n_satz41f : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (n:l_not (l_e_st_eq_landau_n_lessf x y)), l_e_st_eq_landau_n_moreq x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (n:l_not (l_e_st_eq_landau_n_lessf x y)) => l_comor (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_or3_th3 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41a x y) n)))). Time Defined. (* constant 1178 *) Definition l_e_st_eq_landau_n_satz41g : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_not (l_e_st_eq_landau_n_lesseq x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_or_th3 (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) (l_ec3e23 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41b x y) m) (l_ec3e21 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41b x y) m)))). Time Defined. (* constant 1179 *) Definition l_e_st_eq_landau_n_satz41h : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_not (l_e_st_eq_landau_n_moreq x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_or_th3 (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) (l_ec3e32 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41b x y) l) (l_ec3e31 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41b x y) l)))). Time Defined. (* constant 1180 *) Definition l_e_st_eq_landau_n_satz41j : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (n:l_not (l_e_st_eq_landau_n_moreq x y)), l_e_st_eq_landau_n_lessf x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (n:l_not (l_e_st_eq_landau_n_moreq x y)) => l_or3e3 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41a x y) (l_or_th5 (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) n) (l_or_th4 (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) n)))). Time Defined. (* constant 1181 *) Definition l_e_st_eq_landau_n_satz41k : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (n:l_not (l_e_st_eq_landau_n_lesseq x y)), l_e_st_eq_landau_n_moref x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (n:l_not (l_e_st_eq_landau_n_lesseq x y)) => l_or3e2 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_satz41a x y) (l_or_th4 (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) n) (l_or_th5 (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) n)))). Time Defined. (* constant 1182 *) Definition l_e_st_eq_landau_n_246_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), (forall (n:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moreq z u)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => (fun (n:l_e_st_eq_landau_n_moref x y) => l_ori1 (l_e_st_eq_landau_n_moref z u) (l_e_st_eq_landau_n_eq z u) (l_e_st_eq_landau_n_satz44 x y z u n e f))))))))). Time Defined. (* constant 1183 *) Definition l_e_st_eq_landau_n_246_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), (forall (g:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_moreq z u)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => (fun (g:l_e_st_eq_landau_n_eq x y) => l_ori2 (l_e_st_eq_landau_n_moref z u) (l_e_st_eq_landau_n_eq z u) (l_e_st_eq_landau_n_tr3eq z x y u (l_e_st_eq_landau_n_symeq x z e) g f))))))))). Time Defined. (* constant 1184 *) Definition l_e_st_eq_landau_n_satz46 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), l_e_st_eq_landau_n_moreq z u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => l_orapp (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moreq z u) m (fun (t:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_246_t1 x y z u m e f t) (fun (t:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_246_t2 x y z u m e f t)))))))). Time Defined. (* constant 1185 *) Definition l_e_st_eq_landau_n_eqmoreq12 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), (forall (m:l_e_st_eq_landau_n_moreq x z), l_e_st_eq_landau_n_moreq y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => (fun (m:l_e_st_eq_landau_n_moreq x z) => l_e_st_eq_landau_n_satz46 x z y u m e f))))))). Time Defined. (* constant 1186 *) Definition l_e_st_eq_landau_n_eqmoreq1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (m:l_e_st_eq_landau_n_moreq x z), l_e_st_eq_landau_n_moreq y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (m:l_e_st_eq_landau_n_moreq x z) => l_e_st_eq_landau_n_satz46 x z y z m e (l_e_st_eq_landau_n_refeq z)))))). Time Defined. (* constant 1187 *) Definition l_e_st_eq_landau_n_eqmoreq2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (m:l_e_st_eq_landau_n_moreq z x), l_e_st_eq_landau_n_moreq z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (m:l_e_st_eq_landau_n_moreq z x) => l_e_st_eq_landau_n_satz46 z x z y m (l_e_st_eq_landau_n_refeq z) e))))). Time Defined. (* constant 1188 *) Definition l_e_st_eq_landau_n_247_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), (forall (k:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lesseq z u)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => (fun (k:l_e_st_eq_landau_n_lessf x y) => l_ori1 (l_e_st_eq_landau_n_lessf z u) (l_e_st_eq_landau_n_eq z u) (l_e_st_eq_landau_n_satz45 x y z u k e f))))))))). Time Defined. (* constant 1189 *) Definition l_e_st_eq_landau_n_247_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), (forall (g:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_lesseq z u)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => (fun (g:l_e_st_eq_landau_n_eq x y) => l_ori2 (l_e_st_eq_landau_n_lessf z u) (l_e_st_eq_landau_n_eq z u) (l_e_st_eq_landau_n_tr3eq z x y u (l_e_st_eq_landau_n_symeq x z e) g f))))))))). Time Defined. (* constant 1190 *) Definition l_e_st_eq_landau_n_satz47 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (e:l_e_st_eq_landau_n_eq x z), (forall (f:l_e_st_eq_landau_n_eq y u), l_e_st_eq_landau_n_lesseq z u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (e:l_e_st_eq_landau_n_eq x z) => (fun (f:l_e_st_eq_landau_n_eq y u) => l_orapp (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_lesseq z u) l (fun (t:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_247_t1 x y z u l e f t) (fun (t:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_247_t2 x y z u l e f t)))))))). Time Defined. (* constant 1191 *) Definition l_e_st_eq_landau_n_eqlesseq12 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), (forall (l:l_e_st_eq_landau_n_lesseq x z), l_e_st_eq_landau_n_lesseq y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => (fun (l:l_e_st_eq_landau_n_lesseq x z) => l_e_st_eq_landau_n_satz47 x z y u l e f))))))). Time Defined. (* constant 1192 *) Definition l_e_st_eq_landau_n_eqlesseq1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (l:l_e_st_eq_landau_n_lesseq x z), l_e_st_eq_landau_n_lesseq y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (l:l_e_st_eq_landau_n_lesseq x z) => l_e_st_eq_landau_n_satz47 x z y z l e (l_e_st_eq_landau_n_refeq z)))))). Time Defined. (* constant 1193 *) Definition l_e_st_eq_landau_n_eqlesseq2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (l:l_e_st_eq_landau_n_lesseq z x), l_e_st_eq_landau_n_lesseq z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (l:l_e_st_eq_landau_n_lesseq z x) => l_e_st_eq_landau_n_satz47 z x z y l (l_e_st_eq_landau_n_refeq z) e))))). Time Defined. (* constant 1194 *) Definition l_e_st_eq_landau_n_satz48 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), l_e_st_eq_landau_n_lesseq y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => l_or_th9 (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_lessf y x) (l_e_st_eq_landau_n_eq y x) m (fun (t:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz42 x y t) (fun (t:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz38 x y t)))). Time Defined. (* constant 1195 *) Definition l_e_st_eq_landau_n_satz49 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), l_e_st_eq_landau_n_moreq y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => l_or_th9 (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref y x) (l_e_st_eq_landau_n_eq y x) l (fun (t:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz43 x y t) (fun (t:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz38 x y t)))). Time Defined. (* constant 1196 *) Definition l_e_st_eq_landau_n_250_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lessf y z), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lessf y z) => l_e_st_eq_landau_n_satz34a (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) l k))))). Time Defined. (* constant 1197 *) Definition l_e_st_eq_landau_n_250_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lessf y z), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lessf y z) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)))) (l_e_st_eq_landau_n_250_t1 x y z l k)))))). Time Defined. (* constant 1198 *) Definition l_e_st_eq_landau_n_satz50 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lessf y z), l_e_st_eq_landau_n_lessf x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lessf y z) => l_e_st_eq_landau_n_satz33c (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_250_t2 x y z l k)))))). Time Defined. (* constant 1199 *) Definition l_e_st_eq_landau_n_trlessf : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lessf y z), l_e_st_eq_landau_n_lessf x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lessf y z) => l_e_st_eq_landau_n_satz50 x y z l k))))). Time Defined. (* constant 1200 *) Definition l_e_st_eq_landau_n_trmoref : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moref y z), l_e_st_eq_landau_n_moref x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moref y z) => l_e_st_eq_landau_n_satz43 z x (l_e_st_eq_landau_n_satz50 z y x (l_e_st_eq_landau_n_satz42 y z n) (l_e_st_eq_landau_n_satz42 x y m))))))). Time Defined. (* constant 1201 *) Definition l_e_st_eq_landau_n_satz51a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lessf y z), l_e_st_eq_landau_n_lessf x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lessf y z) => l_orapp (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_lessf x z) l (fun (t:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz50 x y z t k) (fun (t:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqlessf1 y x z (l_e_st_eq_landau_n_symeq x y t) k)))))). Time Defined. (* constant 1202 *) Definition l_e_st_eq_landau_n_satz51b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), l_e_st_eq_landau_n_lessf x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => l_orapp (l_e_st_eq_landau_n_lessf y z) (l_e_st_eq_landau_n_eq y z) (l_e_st_eq_landau_n_lessf x z) k (fun (t:l_e_st_eq_landau_n_lessf y z) => l_e_st_eq_landau_n_satz50 x y z l t) (fun (t:l_e_st_eq_landau_n_eq y z) => l_e_st_eq_landau_n_eqlessf2 y z x t l)))))). Time Defined. (* constant 1203 *) Definition l_e_st_eq_landau_n_satz51c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moref y z), l_e_st_eq_landau_n_moref x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moref y z) => l_e_st_eq_landau_n_satz43 z x (l_e_st_eq_landau_n_satz51b z y x (l_e_st_eq_landau_n_satz42 y z n) (l_e_st_eq_landau_n_satz48 x y m))))))). Time Defined. (* constant 1204 *) Definition l_e_st_eq_landau_n_satz51d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moreq y z), l_e_st_eq_landau_n_moref x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moreq y z) => l_e_st_eq_landau_n_satz43 z x (l_e_st_eq_landau_n_satz51a z y x (l_e_st_eq_landau_n_satz48 y z n) (l_e_st_eq_landau_n_satz42 x y m))))))). Time Defined. (* constant 1205 *) Definition l_e_st_eq_landau_n_252_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq y z), l_e_st_eq_landau_n_lesseq x z))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq y z) => l_ori2 (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_eq x z) (l_e_st_eq_landau_n_treq x y z e f)))))))). Time Defined. (* constant 1206 *) Definition l_e_st_eq_landau_n_252_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (j:l_e_st_eq_landau_n_lessf y z), l_e_st_eq_landau_n_lesseq x z))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (j:l_e_st_eq_landau_n_lessf y z) => l_ori1 (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_eq x z) (l_e_st_eq_landau_n_satz51a x y z l j)))))))). Time Defined. (* constant 1207 *) Definition l_e_st_eq_landau_n_252_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_lesseq x z)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_orapp (l_e_st_eq_landau_n_lessf y z) (l_e_st_eq_landau_n_eq y z) (l_e_st_eq_landau_n_lesseq x z) k (fun (t:l_e_st_eq_landau_n_lessf y z) => l_e_st_eq_landau_n_252_t2 x y z l k e t) (fun (t:l_e_st_eq_landau_n_eq y z) => l_e_st_eq_landau_n_252_t1 x y z l k e t))))))). Time Defined. (* constant 1208 *) Definition l_e_st_eq_landau_n_252_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), (forall (j:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lesseq x z)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => (fun (j:l_e_st_eq_landau_n_lessf x y) => l_ori1 (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_eq x z) (l_e_st_eq_landau_n_satz51b x y z j k))))))). Time Defined. (* constant 1209 *) Definition l_e_st_eq_landau_n_satz52 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), l_e_st_eq_landau_n_lesseq x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => l_orapp (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_lesseq x z) l (fun (t:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_252_t4 x y z l k t) (fun (t:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_252_t3 x y z l k t)))))). Time Defined. (* constant 1210 *) Definition l_e_st_eq_landau_n_trlesseq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), l_e_st_eq_landau_n_lesseq x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => l_e_st_eq_landau_n_satz52 x y z l k))))). Time Defined. (* constant 1211 *) Definition l_e_st_eq_landau_n_252_t5 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), (forall (j:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lesseq x z)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => (fun (j:l_e_st_eq_landau_n_lessf x y) => l_ori1 (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_eq x z) (l_e_st_eq_landau_n_satz51b x y z j k))))))). Time Defined. (* constant 1212 *) Definition l_e_st_eq_landau_n_252_t6 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_lesseq x z)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqlesseq1 y x z (l_e_st_eq_landau_n_symeq x y e) k)))))). Time Defined. (* constant 1213 *) Definition l_e_st_eq_landau_n_252_anders : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq y z), l_e_st_eq_landau_n_lesseq x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq y z) => l_orapp (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_lesseq x z) l (fun (t:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_252_t5 x y z l k t) (fun (t:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_252_t6 x y z l k t)))))). Time Defined. (* constant 1214 *) Definition l_e_st_eq_landau_n_trmoreq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq y z), l_e_st_eq_landau_n_moreq x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq y z) => l_e_st_eq_landau_n_satz49 z x (l_e_st_eq_landau_n_satz52 z y x (l_e_st_eq_landau_n_satz48 y z n) (l_e_st_eq_landau_n_satz48 x y m))))))). Time Defined. (* constant 1215 *) Definition l_e_st_eq_landau_n_253_t1 : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x))). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_distpt1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_satz18 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1216 *) Definition l_e_st_eq_landau_n_253_t2 : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_2x x)) x). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_morefi2 x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_253_t1 x)). Time Defined. (* constant 1217 *) Definition l_e_st_eq_landau_n_satz53 : (forall (x:l_e_st_eq_landau_n_frac), l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_moref t x)). exact (fun (x:l_e_st_eq_landau_n_frac) => l_somei l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_moref t x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_253_t2 x)). Time Defined. (* constant 1218 *) Definition l_e_st_eq_landau_n_254_t1 : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2x x)))). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_isless2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_distpt2 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_satz18a (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1219 *) Definition l_e_st_eq_landau_n_254_t2 : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2x x))) x). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_lessfi2 x (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_254_t1 x)). Time Defined. (* constant 1220 *) Definition l_e_st_eq_landau_n_satz54 : (forall (x:l_e_st_eq_landau_n_frac), l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_lessf t x)). exact (fun (x:l_e_st_eq_landau_n_frac) => l_somei l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_lessf t x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_254_t2 x)). Time Defined. (* constant 1221 *) Definition l_e_st_eq_landau_n_255_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz19f (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) l))). Time Defined. (* constant 1222 *) Definition l_e_st_eq_landau_n_255_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz19c (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) l))). Time Defined. (* constant 1223 *) Definition l_e_st_eq_landau_n_255_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_distpt2 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_distpt1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_255_t1 x y l)))). Time Defined. (* constant 1224 *) Definition l_e_st_eq_landau_n_255_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lessf x (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_lessfi1 x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_255_t3 x y l)))). Time Defined. (* constant 1225 *) Definition l_e_st_eq_landau_n_255_t5 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_distpt1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_distpt2 (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_255_t2 x y l)))). Time Defined. (* constant 1226 *) Definition l_e_st_eq_landau_n_255_t6 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_lessfi2 y (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_255_t5 x y l)))). Time Defined. (* constant 1227 *) Definition l_e_st_eq_landau_n_255_t7 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_and (l_e_st_eq_landau_n_lessf x (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_andi (l_e_st_eq_landau_n_lessf x (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) y) (l_e_st_eq_landau_n_255_t4 x y l) (l_e_st_eq_landau_n_255_t6 x y l)))). Time Defined. (* constant 1228 *) Definition l_e_st_eq_landau_n_satz55 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_and (l_e_st_eq_landau_n_lessf x t) (l_e_st_eq_landau_n_lessf t y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_somei l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_and (l_e_st_eq_landau_n_lessf x t) (l_e_st_eq_landau_n_lessf t y)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_255_t7 x y l)))). Time Defined. (* constant 1229 *) Definition l_e_st_eq_landau_n_pf : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_frac)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))). Time Defined. (* constant 1230 *) Definition l_e_st_eq_landau_n_ii3_t1 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_ts y1 x2) (l_e_st_eq_landau_n_ndis12 x1 x2 y1 y2) (l_e_st_eq_landau_n_ndis12 y1 y2 x1 x2))))). Time Defined. (* constant 1231 *) Definition l_e_st_eq_landau_n_ii3_t2 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts x2 y2))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists12 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)) x2 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2)) y2 (l_e_st_eq_landau_n_denis x1 x2) (l_e_st_eq_landau_n_denis y1 y2))))). Time Defined. (* constant 1232 *) Definition l_e_st_eq_landau_n_pf12 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) (l_e_st_eq_landau_n_ts x2 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_isnd (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) (l_e_st_eq_landau_n_ts x2 y2) (l_e_st_eq_landau_n_ii3_t1 x1 x2 y1 y2) (l_e_st_eq_landau_n_ii3_t2 x1 x2 y1 y2))))). Time Defined. (* constant 1233 *) Definition l_e_st_eq_landau_n_ii3_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ndisn2 x n1 n2) (l_e_st_eq_landau_n_ndis1d x n1 n2)))). Time Defined. (* constant 1234 *) Definition l_e_st_eq_landau_n_ii3_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) n2 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_denis n1 n2)))). Time Defined. (* constant 1235 *) Definition l_e_st_eq_landau_n_pf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_isnd (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2) (l_e_st_eq_landau_n_ii3_t3 x n1 n2) (l_e_st_eq_landau_n_ii3_t4 x n1 n2)))). Time Defined. (* constant 1236 *) Definition l_e_st_eq_landau_n_ii3_t5 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ndis1d x n1 n2) (l_e_st_eq_landau_n_ndisn2 x n1 n2)))). Time Defined. (* constant 1237 *) Definition l_e_st_eq_landau_n_ii3_t6 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists1 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) n2 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_denis n1 n2)))). Time Defined. (* constant 1238 *) Definition l_e_st_eq_landau_n_pf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr n1 n2) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_isnd (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ii3_t5 x n1 n2) (l_e_st_eq_landau_n_ii3_t6 x n1 n2)))). Time Defined. (* constant 1239 *) Definition l_e_st_eq_landau_n_pfeq12a : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) (l_e_st_eq_landau_n_ts x2 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) (l_e_st_eq_landau_n_ts x2 y2)) (l_e_st_eq_landau_n_pf12 x1 x2 y1 y2))))). Time Defined. (* constant 1240 *) Definition l_e_st_eq_landau_n_pfeq12b : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) (l_e_st_eq_landau_n_ts x2 y2)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq2 (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 y2) (l_e_st_eq_landau_n_ts y1 x2)) (l_e_st_eq_landau_n_ts x2 y2)) (l_e_st_eq_landau_n_pf12 x1 x2 y1 y2))))). Time Defined. (* constant 1241 *) Definition l_e_st_eq_landau_n_pfeq1a : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2)) (l_e_st_eq_landau_n_pf1 x n1 n2)))). Time Defined. (* constant 1242 *) Definition l_e_st_eq_landau_n_pfeq1b : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2)) (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_fr n1 n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq2 (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2)) (l_e_st_eq_landau_n_pf1 x n1 n2)))). Time Defined. (* constant 1243 *) Definition l_e_st_eq_landau_n_pfeq2a : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr n1 n2) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr n1 n2) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_pf2 x n1 n2)))). Time Defined. (* constant 1244 *) Definition l_e_st_eq_landau_n_pfeq2b : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr n1 n2) x)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq2 (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr n1 n2) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n2)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_pf2 x n1 n2)))). Time Defined. (* constant 1245 *) Definition l_e_st_eq_landau_n_356_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_ists1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u)) e)))))). Time Defined. (* constant 1246 *) Definition l_e_st_eq_landau_n_356_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_356_t1 z u x y f e)))))). Time Defined. (* constant 1247 *) Definition l_e_st_eq_landau_n_356_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2u x y z u) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2u x y z u) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2u x y z u) (l_e_st_eq_landau_n_2z x y z))))))))). Time Defined. (* constant 1248 *) Definition l_e_st_eq_landau_n_356_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_356_t3 x y z u e f) (l_e_st_eq_landau_n_356_t1 x y z u e f) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2x x))))))))). Time Defined. (* constant 1249 *) Definition l_e_st_eq_landau_n_356_t5 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_356_t2 x y z u e f) (l_e_st_eq_landau_n_stets (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))))))). Time Defined. (* constant 1250 *) Definition l_e_st_eq_landau_n_356_t6 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_356_t4 x y z u e f) (l_e_st_eq_landau_n_356_t5 x y z u e f))))))). Time Defined. (* constant 1251 *) Definition l_e_st_eq_landau_n_356_t7 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_disttp1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_356_t6 x y z u e f) (l_e_st_eq_landau_n_distpt1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))))))))). Time Defined. (* constant 1252 *) Definition l_e_st_eq_landau_n_satz56 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_eqi12 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_356_t7 x y z u e f))))))). Time Defined. (* constant 1253 *) Definition l_e_st_eq_landau_n_eqpf12 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_satz56 x y z u e f)))))). Time Defined. (* constant 1254 *) Definition l_e_st_eq_landau_n_eqpf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqpf12 x y z z e (l_e_st_eq_landau_n_refeq z))))). Time Defined. (* constant 1255 *) Definition l_e_st_eq_landau_n_eqpf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqpf12 z z x y (l_e_st_eq_landau_n_refeq z) e)))). Time Defined. (* constant 1256 *) Definition l_e_st_eq_landau_n_satz57 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 n) (l_e_st_eq_landau_n_fr x2 n)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl x1 x2) n)))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 n) (l_e_st_eq_landau_n_fr x2 n)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 n) (l_e_st_eq_landau_n_ts x2 n)) (l_e_st_eq_landau_n_ts n n)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl x1 x2) n) (l_e_st_eq_landau_n_ts n n)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl x1 x2) n) (l_e_st_eq_landau_n_pfeq12a x1 n x2 n) (l_e_st_eq_landau_n_eqn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x1 n) (l_e_st_eq_landau_n_ts x2 n)) (l_e_st_eq_landau_n_ts n n) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl x1 x2) n) (l_e_st_eq_landau_n_distpt1 x1 x2 n)) (l_e_st_eq_landau_n_satz40c (l_e_st_eq_landau_n_pl x1 x2) n n)))). Time Defined. (* constant 1257 *) Definition l_e_st_eq_landau_n_satz57a : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl x1 x2) n) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 n) (l_e_st_eq_landau_n_fr x2 n))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_symeq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x1 n) (l_e_st_eq_landau_n_fr x2 n)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl x1 x2) n) (l_e_st_eq_landau_n_satz57 x1 x2 n)))). Time Defined. (* constant 1258 *) Definition l_e_st_eq_landau_n_satz58 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x y) (l_e_st_eq_landau_n_pf y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eqnd (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_compl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))). Time Defined. (* constant 1259 *) Definition l_e_st_eq_landau_n_compf : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x y) (l_e_st_eq_landau_n_pf y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz58 x y)). Time Defined. (* constant 1260 *) Definition l_e_st_eq_landau_n_359_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ists1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)))))). Time Defined. (* constant 1261 *) Definition l_e_st_eq_landau_n_359_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_359_t1 x y z)))). Time Defined. (* constant 1262 *) Definition l_e_st_eq_landau_n_359_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_disttp1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_359_t2 x y z)))). Time Defined. (* constant 1263 *) Definition l_e_st_eq_landau_n_359_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_assts2 (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x))))). Time Defined. (* constant 1264 *) Definition l_e_st_eq_landau_n_359_t5 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_359_t3 x y z) (l_e_st_eq_landau_n_359_t4 x y z)))). Time Defined. (* constant 1265 *) Definition l_e_st_eq_landau_n_359_t6 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_distpt1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))))). Time Defined. (* constant 1266 *) Definition l_e_st_eq_landau_n_359_t7 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_359_t5 x y z) (l_e_st_eq_landau_n_asspl1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_359_t6 x y z)))). Time Defined. (* constant 1267 *) Definition l_e_st_eq_landau_n_satz59 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_pfeq2a z (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_eqnd (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_359_t7 x y z) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_pfeq1b x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))))). Time Defined. (* constant 1268 *) Definition l_e_st_eq_landau_n_asspf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz59 x y z))). Time Defined. (* constant 1269 *) Definition l_e_st_eq_landau_n_asspf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_pf x y) z)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_symeq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_asspf1 x y z)))). Time Defined. (* constant 1270 *) Definition l_e_st_eq_landau_n_stets1 : (forall (a:l_e_st_eq_landau_n_nat), (forall (b:l_e_st_eq_landau_n_nat), (forall (c:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) c) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a c) b)))). exact (fun (a:l_e_st_eq_landau_n_nat) => (fun (b:l_e_st_eq_landau_n_nat) => (fun (c:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) c) (l_e_st_eq_landau_n_ts a (l_e_st_eq_landau_n_ts b c)) (l_e_st_eq_landau_n_ts a (l_e_st_eq_landau_n_ts c b)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a c) b) (l_e_st_eq_landau_n_assts1 a b c) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts b c) (l_e_st_eq_landau_n_ts c b) a (l_e_st_eq_landau_n_comts b c)) (l_e_st_eq_landau_n_assts2 a c b)))). Time Defined. (* constant 1271 *) Definition l_e_st_eq_landau_n_359_t8 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_disttp1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_stets1 (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)))))). Time Defined. (* constant 1272 *) Definition l_e_st_eq_landau_n_359_t9 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_assts2 (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_stets1 (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))))). Time Defined. (* constant 1273 *) Definition l_e_st_eq_landau_n_359_anderst7 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_359_t8 x y z) (l_e_st_eq_landau_n_359_t9 x y z)) (l_e_st_eq_landau_n_asspl1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_distpt1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)))))). Time Defined. (* constant 1274 *) Definition l_e_st_eq_landau_n_360_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz18 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)))). Time Defined. (* constant 1275 *) Definition l_e_st_eq_landau_n_360_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz32a (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_360_t1 x y))). Time Defined. (* constant 1276 *) Definition l_e_st_eq_landau_n_360_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x))))). Time Defined. (* constant 1277 *) Definition l_e_st_eq_landau_n_360_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_ismore2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_360_t3 x y) (l_e_st_eq_landau_n_360_t2 x y))). Time Defined. (* constant 1278 *) Definition l_e_st_eq_landau_n_satz60 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x y) x)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_morefi2 x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_360_t4 x y))). Time Defined. (* constant 1279 *) Definition l_e_st_eq_landau_n_satz60a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_lessf x (l_e_st_eq_landau_n_pf x y))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_pf x y) x (l_e_st_eq_landau_n_satz60 x y))). Time Defined. (* constant 1280 *) Definition l_e_st_eq_landau_n_361_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz32a (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z) m)))). Time Defined. (* constant 1281 *) Definition l_e_st_eq_landau_n_361_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_stets1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_stets1 (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_361_t1 x y z m))))). Time Defined. (* constant 1282 *) Definition l_e_st_eq_landau_n_361_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_stets1 (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))))). Time Defined. (* constant 1283 *) Definition l_e_st_eq_landau_n_361_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz19h (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_361_t3 x y z m) (l_e_st_eq_landau_n_361_t2 x y z m))))). Time Defined. (* constant 1284 *) Definition l_e_st_eq_landau_n_361_t5 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_distpt1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_distpt1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_361_t4 x y z m))))). Time Defined. (* constant 1285 *) Definition l_e_st_eq_landau_n_361_t6 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz32a (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_361_t5 x y z m))))). Time Defined. (* constant 1286 *) Definition l_e_st_eq_landau_n_361_t7 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_361_t6 x y z m))))). Time Defined. (* constant 1287 *) Definition l_e_st_eq_landau_n_satz61 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_morefi12 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_361_t7 x y z m))))). Time Defined. (* constant 1288 *) Definition l_e_st_eq_landau_n_satz62a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz61 x y z m)))). Time Defined. (* constant 1289 *) Definition l_e_st_eq_landau_n_satz62b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqpf1 x y z e)))). Time Defined. (* constant 1290 *) Definition l_e_st_eq_landau_n_satz62c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_satz61 y x z (l_e_st_eq_landau_n_satz43 x y l)))))). Time Defined. (* constant 1291 *) Definition l_e_st_eq_landau_n_satz62d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_eqmoref12 (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_pf z y) (l_e_st_eq_landau_n_compf x z) (l_e_st_eq_landau_n_compf y z) (l_e_st_eq_landau_n_satz62a x y z m))))). Time Defined. (* constant 1292 *) Definition l_e_st_eq_landau_n_satz62e : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqpf2 x y z e)))). Time Defined. (* constant 1293 *) Definition l_e_st_eq_landau_n_satz62f : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_eqlessf12 (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_pf z y) (l_e_st_eq_landau_n_compf x z) (l_e_st_eq_landau_n_compf y z) (l_e_st_eq_landau_n_satz62c x y z l))))). Time Defined. (* constant 1294 *) Definition l_e_st_eq_landau_n_satz62g : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (m:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (m:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_eqmoref2 (l_e_st_eq_landau_n_pf x u) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_eqpf1 x y u e) (l_e_st_eq_landau_n_satz62d z u x m))))))). Time Defined. (* constant 1295 *) Definition l_e_st_eq_landau_n_satz62h : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (m:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf u y))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (m:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_eqmoref12 (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf u y) (l_e_st_eq_landau_n_compf x z) (l_e_st_eq_landau_n_compf y u) (l_e_st_eq_landau_n_satz62g x y z u e m))))))). Time Defined. (* constant 1296 *) Definition l_e_st_eq_landau_n_satz62j : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (l:l_e_st_eq_landau_n_lessf z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (l:l_e_st_eq_landau_n_lessf z u) => l_e_st_eq_landau_n_eqlessf2 (l_e_st_eq_landau_n_pf x u) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_eqpf1 x y u e) (l_e_st_eq_landau_n_satz62f z u x l))))))). Time Defined. (* constant 1297 *) Definition l_e_st_eq_landau_n_satz62k : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (l:l_e_st_eq_landau_n_lessf z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf u y))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (l:l_e_st_eq_landau_n_lessf z u) => l_e_st_eq_landau_n_eqlessf12 (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf u y) (l_e_st_eq_landau_n_compf x z) (l_e_st_eq_landau_n_compf y u) (l_e_st_eq_landau_n_satz62j x y z u e l))))))). Time Defined. (* constant 1298 *) Definition l_e_st_eq_landau_n_363_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_or3 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz41a x y))). Time Defined. (* constant 1299 *) Definition l_e_st_eq_landau_n_363_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_ec3 (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz41b (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)))). Time Defined. (* constant 1300 *) Definition l_e_st_eq_landau_n_satz63a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)), l_e_st_eq_landau_n_moref x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) => l_ec3_th11 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_363_t1 x y z) (l_e_st_eq_landau_n_363_t2 x y z) (fun (u:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz62b x y z u) (fun (u:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz62a x y z u) (fun (u:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz62c x y z u) m)))). Time Defined. (* constant 1301 *) Definition l_e_st_eq_landau_n_satz63b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)), l_e_st_eq_landau_n_eq x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) => l_ec3_th10 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_363_t1 x y z) (l_e_st_eq_landau_n_363_t2 x y z) (fun (u:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz62b x y z u) (fun (u:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz62a x y z u) (fun (u:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz62c x y z u) e)))). Time Defined. (* constant 1302 *) Definition l_e_st_eq_landau_n_satz63c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)), l_e_st_eq_landau_n_lessf x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) => l_ec3_th12 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_363_t1 x y z) (l_e_st_eq_landau_n_363_t2 x y z) (fun (u:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz62b x y z u) (fun (u:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz62a x y z u) (fun (u:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz62c x y z u) l)))). Time Defined. (* constant 1303 *) Definition l_e_st_eq_landau_n_satz63d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y)), l_e_st_eq_landau_n_moref x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y)) => l_e_st_eq_landau_n_satz63a x y z (l_e_st_eq_landau_n_eqmoref12 (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf z y) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_compf z x) (l_e_st_eq_landau_n_compf z y) m))))). Time Defined. (* constant 1304 *) Definition l_e_st_eq_landau_n_satz63e : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y)), l_e_st_eq_landau_n_eq x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y)) => l_e_st_eq_landau_n_satz63b x y z (l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_compf x z) e (l_e_st_eq_landau_n_compf z y)))))). Time Defined. (* constant 1305 *) Definition l_e_st_eq_landau_n_satz63f : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (f:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y)), l_e_st_eq_landau_n_lessf x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (f:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf z y)) => l_e_st_eq_landau_n_satz63c x y z (l_e_st_eq_landau_n_eqlessf12 (l_e_st_eq_landau_n_pf z x) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf z y) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_compf z x) (l_e_st_eq_landau_n_compf z y) f))))). Time Defined. (* constant 1306 *) Definition l_e_st_eq_landau_n_364_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_satz61 x y z m)))))). Time Defined. (* constant 1307 *) Definition l_e_st_eq_landau_n_364_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_eqmoref12 (l_e_st_eq_landau_n_pf z y) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_pf u y) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_compf z y) (l_e_st_eq_landau_n_compf u y) (l_e_st_eq_landau_n_satz61 z u y n))))))). Time Defined. (* constant 1308 *) Definition l_e_st_eq_landau_n_satz64 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_trmoref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_364_t1 x y z u m n) (l_e_st_eq_landau_n_364_t2 x y z u m n))))))). Time Defined. (* constant 1309 *) Definition l_e_st_eq_landau_n_satz64a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lessf z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lessf z u) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_satz64 y x u z (l_e_st_eq_landau_n_satz43 x y l) (l_e_st_eq_landau_n_satz43 z u k)))))))). Time Defined. (* constant 1310 *) Definition l_e_st_eq_landau_n_satz65a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moref z u) => l_orapp (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u)) m (fun (v:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz64 x y z u v n) (fun (v:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz62g x y z u v n))))))). Time Defined. (* constant 1311 *) Definition l_e_st_eq_landau_n_satz65b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moreq z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => l_orapp (l_e_st_eq_landau_n_moref z u) (l_e_st_eq_landau_n_eq z u) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u)) n (fun (v:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_satz64 x y z u m v) (fun (v:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_eqmoref2 (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_eqpf2 z u y v) (l_e_st_eq_landau_n_satz61 x y z m)))))))). Time Defined. (* constant 1312 *) Definition l_e_st_eq_landau_n_satz65c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lessf z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lessf z u) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_satz65a y x u z (l_e_st_eq_landau_n_satz49 x y l) (l_e_st_eq_landau_n_satz43 z u k)))))))). Time Defined. (* constant 1313 *) Definition l_e_st_eq_landau_n_satz65d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lesseq z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lesseq z u) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_satz65b y x u z (l_e_st_eq_landau_n_satz43 x y l) (l_e_st_eq_landau_n_satz49 z u k)))))))). Time Defined. (* constant 1314 *) Definition l_e_st_eq_landau_n_366_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_moreqi2 (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_satz56 x y z u e f))))))))). Time Defined. (* constant 1315 *) Definition l_e_st_eq_landau_n_366_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (o:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (o:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_moreqi1 (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_satz65a x y z u m o))))))))). Time Defined. (* constant 1316 *) Definition l_e_st_eq_landau_n_366_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_orapp (l_e_st_eq_landau_n_moref z u) (l_e_st_eq_landau_n_eq z u) (l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u)) n (fun (v:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_366_t2 x y z u m n e v) (fun (v:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_366_t1 x y z u m n e v)))))))). Time Defined. (* constant 1317 *) Definition l_e_st_eq_landau_n_366_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), (forall (o:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => (fun (o:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_moreqi1 (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_satz65b x y z u o n)))))))). Time Defined. (* constant 1318 *) Definition l_e_st_eq_landau_n_satz66 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => l_orapp (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u)) m (fun (v:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_366_t4 x y z u m n v) (fun (v:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_366_t3 x y z u m n v))))))). Time Defined. (* constant 1319 *) Definition l_e_st_eq_landau_n_satz66a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq z u), l_e_st_eq_landau_n_lesseq (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq z u) => l_e_st_eq_landau_n_satz48 (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_satz66 y x u z (l_e_st_eq_landau_n_satz49 x y l) (l_e_st_eq_landau_n_satz49 z u k)))))))). Time Defined. (* constant 1320 *) Definition l_e_st_eq_landau_n_367_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (v:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x), l_e_st_eq_landau_n_moref x y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x) => l_e_st_eq_landau_n_eqmoref1 (l_e_st_eq_landau_n_pf y v) x y e (l_e_st_eq_landau_n_satz60 y v)))))). Time Defined. (* constant 1321 *) Definition l_e_st_eq_landau_n_367_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (v:l_e_st_eq_landau_n_frac), l_not (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (v:l_e_st_eq_landau_n_frac) => l_imp_th3 (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_satz41d x y l) (fun (t:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x) => l_e_st_eq_landau_n_367_t1 x y l v t))))). Time Defined. (* constant 1322 *) Definition l_e_st_eq_landau_n_vorbemerkung67 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), l_not (l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y t) x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => l_some_th5 l_e_st_eq_landau_n_frac (fun (v:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x) (fun (v:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_367_t2 x y l v)))). Time Defined. (* constant 1323 *) Definition l_e_st_eq_landau_n_satz67b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (v:l_e_st_eq_landau_n_frac), (forall (w:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x), (forall (f:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y w) x), l_e_st_eq_landau_n_eq v w)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (w:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x) => (fun (f:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y w) x) => l_e_st_eq_landau_n_satz63e v w y (l_e_st_eq_landau_n_treq2 (l_e_st_eq_landau_n_pf y v) (l_e_st_eq_landau_n_pf y w) x e f))))))). Time Defined. (* constant 1324 *) Definition l_e_st_eq_landau_n_367_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_one (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) t)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_onei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) t) (l_e_st_eq_landau_n_satz8b (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x))) m))). Time Defined. (* constant 1325 *) Definition l_e_st_eq_landau_n_367_vo : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_nat))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_ind l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) t) (l_e_st_eq_landau_n_367_t3 x y m)))). Time Defined. (* constant 1326 *) Definition l_e_st_eq_landau_n_367_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_367_vo x y m))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_oneax l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) t) (l_e_st_eq_landau_n_367_t3 x y m)))). Time Defined. (* constant 1327 *) Definition l_e_st_eq_landau_n_367_w : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_frac))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_367_vo x y m) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))))). Time Defined. (* constant 1328 *) Definition l_e_st_eq_landau_n_367_t5 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_eq y (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_treq y (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_satz40 y (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_eqd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x)))))). Time Defined. (* constant 1329 *) Definition l_e_st_eq_landau_n_367_t6 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y (l_e_st_eq_landau_n_367_w x y m)) x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_tr4eq (l_e_st_eq_landau_n_pf y (l_e_st_eq_landau_n_367_w x y m)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_367_vo x y m) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_367_vo x y m)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) x (l_e_st_eq_landau_n_eqpf1 y (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_367_w x y m) (l_e_st_eq_landau_n_367_t5 x y m)) (l_e_st_eq_landau_n_satz57 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_367_vo x y m) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_eqn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_367_vo x y m)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_367_vo x y m)) (l_e_st_eq_landau_n_367_t4 x y m))) (l_e_st_eq_landau_n_satz40a x (l_e_st_eq_landau_n_2y x y))))). Time Defined. (* constant 1330 *) Definition l_e_st_eq_landau_n_satz67a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y t) x)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_somei l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y t) x) (l_e_st_eq_landau_n_367_w x y m) (l_e_st_eq_landau_n_367_t6 x y m)))). Time Defined. (* constant 1331 *) Definition l_e_st_eq_landau_n_mf : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_frac))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_367_w x y m))). Time Defined. (* constant 1332 *) Definition l_e_st_eq_landau_n_satz67c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y (l_e_st_eq_landau_n_mf x y m)) x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_367_t6 x y m))). Time Defined. (* constant 1333 *) Definition l_e_st_eq_landau_n_satz67d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_eq x (l_e_st_eq_landau_n_pf y (l_e_st_eq_landau_n_mf x y m))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_symeq (l_e_st_eq_landau_n_pf y (l_e_st_eq_landau_n_mf x y m)) x (l_e_st_eq_landau_n_satz67c x y m)))). Time Defined. (* constant 1334 *) Definition l_e_st_eq_landau_n_satz67e : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (v:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x), l_e_st_eq_landau_n_eq v (l_e_st_eq_landau_n_mf x y m)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x) => l_e_st_eq_landau_n_satz67b x y v (l_e_st_eq_landau_n_mf x y m) e (l_e_st_eq_landau_n_satz67c x y m)))))). Time Defined. (* constant 1335 *) Definition l_e_st_eq_landau_n_tf : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_frac)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))). Time Defined. (* constant 1336 *) Definition l_e_st_eq_landau_n_ii4_t1 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts x1 y1))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists12 (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) x1 (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2)) y1 (l_e_st_eq_landau_n_numis x1 x2) (l_e_st_eq_landau_n_numis y1 y2))))). Time Defined. (* constant 1337 *) Definition l_e_st_eq_landau_n_ii4_t2 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts x2 y2))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists12 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)) x2 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2)) y2 (l_e_st_eq_landau_n_denis x1 x2) (l_e_st_eq_landau_n_denis y1 y2))))). Time Defined. (* constant 1338 *) Definition l_e_st_eq_landau_n_tf12 : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x1 y1) (l_e_st_eq_landau_n_ts x2 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_isnd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y1 y2))) (l_e_st_eq_landau_n_ts x1 y1) (l_e_st_eq_landau_n_ts x2 y2) (l_e_st_eq_landau_n_ii4_t1 x1 x2 y1 y2) (l_e_st_eq_landau_n_ii4_t2 x1 x2 y1 y2))))). Time Defined. (* constant 1339 *) Definition l_e_st_eq_landau_n_ii4_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n1)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) n1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_numis n1 n2)))). Time Defined. (* constant 1340 *) Definition l_e_st_eq_landau_n_ii4_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) n2 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_denis n1 n2)))). Time Defined. (* constant 1341 *) Definition l_e_st_eq_landau_n_tf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_isnd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2) (l_e_st_eq_landau_n_ii4_t3 x n1 n2) (l_e_st_eq_landau_n_ii4_t4 x n1 n2)))). Time Defined. (* constant 1342 *) Definition l_e_st_eq_landau_n_ii4_t5 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_1x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists1 (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) n1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_numis n1 n2)))). Time Defined. (* constant 1343 *) Definition l_e_st_eq_landau_n_ii4_t6 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_ists1 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) n2 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_denis n1 n2)))). Time Defined. (* constant 1344 *) Definition l_e_st_eq_landau_n_tf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_is l_e_st_eq_landau_n_frac (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr n1 n2) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_isnd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ii4_t5 x n1 n2) (l_e_st_eq_landau_n_ii4_t6 x n1 n2)))). Time Defined. (* constant 1345 *) Definition l_e_st_eq_landau_n_tfeq12a : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x1 y1) (l_e_st_eq_landau_n_ts x2 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x1 y1) (l_e_st_eq_landau_n_ts x2 y2)) (l_e_st_eq_landau_n_tf12 x1 x2 y1 y2))))). Time Defined. (* constant 1346 *) Definition l_e_st_eq_landau_n_tfeq12b : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), (forall (y1:l_e_st_eq_landau_n_nat), (forall (y2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x1 y1) (l_e_st_eq_landau_n_ts x2 y2)) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)))))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => (fun (y1:l_e_st_eq_landau_n_nat) => (fun (y2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq2 (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x1 x2) (l_e_st_eq_landau_n_fr y1 y2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x1 y1) (l_e_st_eq_landau_n_ts x2 y2)) (l_e_st_eq_landau_n_tf12 x1 x2 y1 y2))))). Time Defined. (* constant 1347 *) Definition l_e_st_eq_landau_n_tfeq1a : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2)) (l_e_st_eq_landau_n_tf1 x n1 n2)))). Time Defined. (* constant 1348 *) Definition l_e_st_eq_landau_n_tfeq1b : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2)) (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_fr n1 n2))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq2 (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_fr n1 n2)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) n1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) n2)) (l_e_st_eq_landau_n_tf1 x n1 n2)))). Time Defined. (* constant 1349 *) Definition l_e_st_eq_landau_n_tfeq2a : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr n1 n2) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr n1 n2) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_tf2 x n1 n2)))). Time Defined. (* constant 1350 *) Definition l_e_st_eq_landau_n_tfeq2b : (forall (x:l_e_st_eq_landau_n_frac), (forall (n1:l_e_st_eq_landau_n_nat), (forall (n2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr n1 n2) x)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (n1:l_e_st_eq_landau_n_nat) => (fun (n2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_refeq2 (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr n1 n2) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts n1 (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts n2 (l_e_st_eq_landau_n_2x x))) (l_e_st_eq_landau_n_tf2 x n1 n2)))). Time Defined. (* constant 1351 *) Definition l_e_st_eq_landau_n_468_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_ists12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)) e f)))))). Time Defined. (* constant 1352 *) Definition l_e_st_eq_landau_n_stets2 : (forall (a:l_e_st_eq_landau_n_nat), (forall (b:l_e_st_eq_landau_n_nat), (forall (c:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a c) (l_e_st_eq_landau_n_ts b d)))))). exact (fun (a:l_e_st_eq_landau_n_nat) => (fun (b:l_e_st_eq_landau_n_nat) => (fun (c:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) (l_e_st_eq_landau_n_ts c d)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a b) (l_e_st_eq_landau_n_ts d c)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a c) (l_e_st_eq_landau_n_ts d b)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts a c) (l_e_st_eq_landau_n_ts b d)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts c d) (l_e_st_eq_landau_n_ts d c) (l_e_st_eq_landau_n_ts a b) (l_e_st_eq_landau_n_comts c d)) (l_e_st_eq_landau_n_stets a b d c) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts d b) (l_e_st_eq_landau_n_ts b d) (l_e_st_eq_landau_n_ts a c) (l_e_st_eq_landau_n_comts d b)))))). Time Defined. (* constant 1353 *) Definition l_e_st_eq_landau_n_468_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2u x y z u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_stets2 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_468_t1 x y z u e f) (l_e_st_eq_landau_n_stets2 (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1u x y z u) (l_e_st_eq_landau_n_2z x y z)))))))). Time Defined. (* constant 1354 *) Definition l_e_st_eq_landau_n_satz68 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_eqi12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1u x y z u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2u x y z u)) (l_e_st_eq_landau_n_468_t2 x y z u e f))))))). Time Defined. (* constant 1355 *) Definition l_e_st_eq_landau_n_eqtf12 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_satz68 x y z u e f)))))). Time Defined. (* constant 1356 *) Definition l_e_st_eq_landau_n_eqtf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqtf12 x y z z e (l_e_st_eq_landau_n_refeq z))))). Time Defined. (* constant 1357 *) Definition l_e_st_eq_landau_n_eqtf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqtf12 z z x y (l_e_st_eq_landau_n_refeq z) e)))). Time Defined. (* constant 1358 *) Definition l_e_st_eq_landau_n_satz69 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eqnd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)))). Time Defined. (* constant 1359 *) Definition l_e_st_eq_landau_n_comtf : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz69 x y)). Time Defined. (* constant 1360 *) Definition l_e_st_eq_landau_n_satz70 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_tf x y) z) (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_tf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_tf x y) z) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_tfeq2a z (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_eqnd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_tfeq1b x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))))). Time Defined. (* constant 1361 *) Definition l_e_st_eq_landau_n_asstf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_tf x y) z) (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_tf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz70 x y z))). Time Defined. (* constant 1362 *) Definition l_e_st_eq_landau_n_asstf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_tf x y) z)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_symeq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_tf x y) z) (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_asstf1 x y z)))). Time Defined. (* constant 1363 *) Definition l_e_st_eq_landau_n_471_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))))) (l_e_st_eq_landau_n_tfeq1a x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_eqn (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_disttp2 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_satz57a (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))))))). Time Defined. (* constant 1364 *) Definition l_e_st_eq_landau_n_471_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_tf x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_treq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_eqnd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_assts2 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_assts2 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_satz40c (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2z x y z))))). Time Defined. (* constant 1365 *) Definition l_e_st_eq_landau_n_471_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_tf x z)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_treq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2y x y)))) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_eqd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2z x y z) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_471_t2 x z y)))). Time Defined. (* constant 1366 *) Definition l_e_st_eq_landau_n_satz71 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_treq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))))) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf x z)) (l_e_st_eq_landau_n_471_t1 x y z) (l_e_st_eq_landau_n_eqpf12 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)))) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_471_t2 x y z) (l_e_st_eq_landau_n_471_t3 x y z))))). Time Defined. (* constant 1367 *) Definition l_e_st_eq_landau_n_disttpf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_tf z (l_e_st_eq_landau_n_pf x y)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_comtf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_satz71 z x y) (l_e_st_eq_landau_n_eqpf12 (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf z y) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_comtf z x) (l_e_st_eq_landau_n_comtf z y))))). Time Defined. (* constant 1368 *) Definition l_e_st_eq_landau_n_disttpf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz71 x y z))). Time Defined. (* constant 1369 *) Definition l_e_st_eq_landau_n_distptf1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_pf x y) z)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_symeq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_disttpf1 x y z)))). Time Defined. (* constant 1370 *) Definition l_e_st_eq_landau_n_distptf2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf x z)) (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_symeq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf x z)) (l_e_st_eq_landau_n_disttpf2 x y z)))). Time Defined. (* constant 1371 *) Definition l_e_st_eq_landau_n_472_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2z x y z))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz32a (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2z x y z)) m)))). Time Defined. (* constant 1372 *) Definition l_e_st_eq_landau_n_472_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z))) (l_e_st_eq_landau_n_stets2 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_stets2 (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1z x y z) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_472_t1 x y z m))))). Time Defined. (* constant 1373 *) Definition l_e_st_eq_landau_n_satz72a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_morefi12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_1z x y z)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_2z x y z)) (l_e_st_eq_landau_n_472_t2 x y z m))))). Time Defined. (* constant 1374 *) Definition l_e_st_eq_landau_n_satz72b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz68 x y z z e (l_e_st_eq_landau_n_refeq z))))). Time Defined. (* constant 1375 *) Definition l_e_st_eq_landau_n_satz72c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_satz72a y x z (l_e_st_eq_landau_n_satz43 x y l)))))). Time Defined. (* constant 1376 *) Definition l_e_st_eq_landau_n_satz72d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_eqmoref12 (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_tf z y) (l_e_st_eq_landau_n_comtf x z) (l_e_st_eq_landau_n_comtf y z) (l_e_st_eq_landau_n_satz72a x y z m))))). Time Defined. (* constant 1377 *) Definition l_e_st_eq_landau_n_satz72e : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_eqtf2 x y z e)))). Time Defined. (* constant 1378 *) Definition l_e_st_eq_landau_n_satz72f : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_eqlessf12 (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_tf z y) (l_e_st_eq_landau_n_comtf x z) (l_e_st_eq_landau_n_comtf y z) (l_e_st_eq_landau_n_satz72c x y z l))))). Time Defined. (* constant 1379 *) Definition l_e_st_eq_landau_n_satz72g : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (m:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (m:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_eqmoref2 (l_e_st_eq_landau_n_tf x u) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_eqtf1 x y u e) (l_e_st_eq_landau_n_satz72d z u x m))))))). Time Defined. (* constant 1380 *) Definition l_e_st_eq_landau_n_satz72h : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (m:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf u y))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (m:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_eqmoref12 (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf u y) (l_e_st_eq_landau_n_comtf x z) (l_e_st_eq_landau_n_comtf y u) (l_e_st_eq_landau_n_satz72g x y z u e m))))))). Time Defined. (* constant 1381 *) Definition l_e_st_eq_landau_n_satz72j : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (l:l_e_st_eq_landau_n_lessf z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (l:l_e_st_eq_landau_n_lessf z u) => l_e_st_eq_landau_n_eqlessf2 (l_e_st_eq_landau_n_tf x u) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_eqtf1 x y u e) (l_e_st_eq_landau_n_satz72f z u x l))))))). Time Defined. (* constant 1382 *) Definition l_e_st_eq_landau_n_satz72k : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (l:l_e_st_eq_landau_n_lessf z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf u y))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (l:l_e_st_eq_landau_n_lessf z u) => l_e_st_eq_landau_n_eqlessf12 (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf u y) (l_e_st_eq_landau_n_comtf x z) (l_e_st_eq_landau_n_comtf y u) (l_e_st_eq_landau_n_satz72j x y z u e l))))))). Time Defined. (* constant 1383 *) Definition l_e_st_eq_landau_n_473_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_or3 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz41a x y))). Time Defined. (* constant 1384 *) Definition l_e_st_eq_landau_n_473_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), l_ec3 (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz41b (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)))). Time Defined. (* constant 1385 *) Definition l_e_st_eq_landau_n_satz73a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)), l_e_st_eq_landau_n_moref x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) => l_ec3_th11 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_473_t1 x y z) (l_e_st_eq_landau_n_473_t2 x y z) (fun (u:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz72b x y z u) (fun (u:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz72a x y z u) (fun (u:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz72c x y z u) m)))). Time Defined. (* constant 1386 *) Definition l_e_st_eq_landau_n_satz73b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)), l_e_st_eq_landau_n_eq x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) => l_ec3_th10 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_473_t1 x y z) (l_e_st_eq_landau_n_473_t2 x y z) (fun (u:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz72b x y z u) (fun (u:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz72a x y z u) (fun (u:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz72c x y z u) e)))). Time Defined. (* constant 1387 *) Definition l_e_st_eq_landau_n_satz73c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)), l_e_st_eq_landau_n_lessf x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) => l_ec3_th12 (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_lessf x y) (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_473_t1 x y z) (l_e_st_eq_landau_n_473_t2 x y z) (fun (u:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz72b x y z u) (fun (u:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz72a x y z u) (fun (u:l_e_st_eq_landau_n_lessf x y) => l_e_st_eq_landau_n_satz72c x y z u) l)))). Time Defined. (* constant 1388 *) Definition l_e_st_eq_landau_n_satz73d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y)), l_e_st_eq_landau_n_moref x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y)) => l_e_st_eq_landau_n_satz73a x y z (l_e_st_eq_landau_n_eqmoref12 (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf z y) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_comtf z x) (l_e_st_eq_landau_n_comtf z y) m))))). Time Defined. (* constant 1389 *) Definition l_e_st_eq_landau_n_satz73e : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y)), l_e_st_eq_landau_n_eq x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y)) => l_e_st_eq_landau_n_satz73b x y z (l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_comtf x z) e (l_e_st_eq_landau_n_comtf z y)))))). Time Defined. (* constant 1390 *) Definition l_e_st_eq_landau_n_satz73f : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y)), l_e_st_eq_landau_n_lessf x y)))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf z y)) => l_e_st_eq_landau_n_satz73c x y z (l_e_st_eq_landau_n_eqlessf12 (l_e_st_eq_landau_n_tf z x) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf z y) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_comtf z x) (l_e_st_eq_landau_n_comtf z y) l))))). Time Defined. (* constant 1391 *) Definition l_e_st_eq_landau_n_474_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_satz72a x y z m)))))). Time Defined. (* constant 1392 *) Definition l_e_st_eq_landau_n_474_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_eqmoref12 (l_e_st_eq_landau_n_tf z y) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_tf u y) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_comtf z y) (l_e_st_eq_landau_n_comtf u y) (l_e_st_eq_landau_n_satz72a z u y n))))))). Time Defined. (* constant 1393 *) Definition l_e_st_eq_landau_n_satz74 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_trmoref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_474_t1 x y z u m n) (l_e_st_eq_landau_n_474_t2 x y z u m n))))))). Time Defined. (* constant 1394 *) Definition l_e_st_eq_landau_n_satz74a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lessf z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lessf z u) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_satz74 y x u z (l_e_st_eq_landau_n_satz43 x y l) (l_e_st_eq_landau_n_satz43 z u k)))))))). Time Defined. (* constant 1395 *) Definition l_e_st_eq_landau_n_satz75a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moref z u) => l_orapp (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u)) m (fun (v:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_satz74 x y z u v n) (fun (v:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_satz72g x y z u v n))))))). Time Defined. (* constant 1396 *) Definition l_e_st_eq_landau_n_satz75b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref x y), (forall (n:l_e_st_eq_landau_n_moreq z u), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => l_orapp (l_e_st_eq_landau_n_moref z u) (l_e_st_eq_landau_n_eq z u) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u)) n (fun (v:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_satz74 x y z u m v) (fun (v:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_eqmoref2 (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_eqtf2 z u y v) (l_e_st_eq_landau_n_satz72a x y z m)))))))). Time Defined. (* constant 1397 *) Definition l_e_st_eq_landau_n_satz75c : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lessf z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lessf z u) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_satz75a y x u z (l_e_st_eq_landau_n_satz49 x y l) (l_e_st_eq_landau_n_satz43 z u k)))))))). Time Defined. (* constant 1398 *) Definition l_e_st_eq_landau_n_satz75d : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf x y), (forall (k:l_e_st_eq_landau_n_lesseq z u), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf x y) => (fun (k:l_e_st_eq_landau_n_lesseq z u) => l_e_st_eq_landau_n_satz42 (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_satz75b y x u z (l_e_st_eq_landau_n_satz43 x y l) (l_e_st_eq_landau_n_satz49 z u k)))))))). Time Defined. (* constant 1399 *) Definition l_e_st_eq_landau_n_476_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_moreqi2 (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_satz68 x y z u e f))))))))). Time Defined. (* constant 1400 *) Definition l_e_st_eq_landau_n_476_t2 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (o:l_e_st_eq_landau_n_moref z u), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (o:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_moreqi1 (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_satz75a x y z u m o))))))))). Time Defined. (* constant 1401 *) Definition l_e_st_eq_landau_n_476_t3 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_orapp (l_e_st_eq_landau_n_moref z u) (l_e_st_eq_landau_n_eq z u) (l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u)) n (fun (v:l_e_st_eq_landau_n_moref z u) => l_e_st_eq_landau_n_476_t2 x y z u m n e v) (fun (v:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_476_t1 x y z u m n e v)))))))). Time Defined. (* constant 1402 *) Definition l_e_st_eq_landau_n_476_t4 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), (forall (o:l_e_st_eq_landau_n_moref x y), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u)))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => (fun (o:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_moreqi1 (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_satz75b x y z u o n)))))))). Time Defined. (* constant 1403 *) Definition l_e_st_eq_landau_n_satz76 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moreq x y), (forall (n:l_e_st_eq_landau_n_moreq z u), l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moreq x y) => (fun (n:l_e_st_eq_landau_n_moreq z u) => l_orapp (l_e_st_eq_landau_n_moref x y) (l_e_st_eq_landau_n_eq x y) (l_e_st_eq_landau_n_moreq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u)) m (fun (v:l_e_st_eq_landau_n_moref x y) => l_e_st_eq_landau_n_476_t4 x y z u m n v) (fun (v:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_landau_n_476_t3 x y z u m n v))))))). Time Defined. (* constant 1404 *) Definition l_e_st_eq_landau_n_satz76a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lesseq x y), (forall (k:l_e_st_eq_landau_n_lesseq z u), l_e_st_eq_landau_n_lesseq (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lesseq x y) => (fun (k:l_e_st_eq_landau_n_lesseq z u) => l_e_st_eq_landau_n_satz48 (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_satz76 y x u z (l_e_st_eq_landau_n_satz49 x y l) (l_e_st_eq_landau_n_satz49 z u k)))))))). Time Defined. (* constant 1405 *) Definition l_e_st_eq_landau_n_satz77b : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (v:l_e_st_eq_landau_n_frac), (forall (w:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y v) x), (forall (f:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y w) x), l_e_st_eq_landau_n_eq v w)))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (w:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y v) x) => (fun (f:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y w) x) => l_e_st_eq_landau_n_satz73e v w y (l_e_st_eq_landau_n_treq2 (l_e_st_eq_landau_n_tf y v) (l_e_st_eq_landau_n_tf y w) x e f))))))). Time Defined. (* constant 1406 *) Definition l_e_st_eq_landau_n_477_v : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_frac)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1y x y)))). Time Defined. (* constant 1407 *) Definition l_e_st_eq_landau_n_477_t1 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y (l_e_st_eq_landau_n_477_v x y)) x)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_tr4eq (l_e_st_eq_landau_n_tf y (l_e_st_eq_landau_n_477_v x y)) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_477_v x y) y) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)))) x (l_e_st_eq_landau_n_comtf y (l_e_st_eq_landau_n_477_v x y)) (l_e_st_eq_landau_n_tfeq2a y (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1y x y))) (l_e_st_eq_landau_n_eqnd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_1y x y))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_1y x y)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y)) (l_e_st_eq_landau_n_1x x) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_2y x y) (l_e_st_eq_landau_n_1y x y)))) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_2x x) (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))) (l_e_st_eq_landau_n_satz40a x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_1y x y) (l_e_st_eq_landau_n_2y x y))))). Time Defined. (* constant 1408 *) Definition l_e_st_eq_landau_n_satz77a : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y t) x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_somei l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y t) x) (l_e_st_eq_landau_n_477_v x y) (l_e_st_eq_landau_n_477_t1 x y))). Time Defined. (* constant 1409 *) Definition l_e_st_eq_landau_n_rt_eq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), Prop)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq x y)). Time Defined. (* constant 1410 *) Definition l_e_st_eq_landau_n_rt_refeq : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_rt_eq x x). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_refeq x). Time Defined. (* constant 1411 *) Definition l_e_st_eq_landau_n_rt_symeq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (t:l_e_st_eq_landau_n_rt_eq x y), l_e_st_eq_landau_n_rt_eq y x))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (t:l_e_st_eq_landau_n_rt_eq x y) => l_e_st_eq_landau_n_symeq x y t))). Time Defined. (* constant 1412 *) Definition l_e_st_eq_landau_n_rt_treq : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (t:l_e_st_eq_landau_n_rt_eq x y), (forall (u:l_e_st_eq_landau_n_rt_eq y z), l_e_st_eq_landau_n_rt_eq x z))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (t:l_e_st_eq_landau_n_rt_eq x y) => (fun (u:l_e_st_eq_landau_n_rt_eq y z) => l_e_st_eq_landau_n_treq x y z t u))))). Time Defined. (* constant 1413 *) Definition l_e_st_eq_landau_n_rt_inf : (forall (x:l_e_st_eq_landau_n_frac), (forall (s:l_e_st_set l_e_st_eq_landau_n_frac), Prop)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (s:l_e_st_set l_e_st_eq_landau_n_frac) => l_e_st_esti l_e_st_eq_landau_n_frac x s)). Time Defined. (* constant 1414 *) Definition l_e_st_eq_landau_n_rt_rat : Type. exact (l_e_st_eq_ect l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq). Time Defined. (* constant 1415 *) Definition l_e_st_eq_landau_n_rt_is : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_is l_e_st_eq_landau_n_rt_rat x0 y0)). Time Defined. (* constant 1416 *) Definition l_e_st_eq_landau_n_rt_nis : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_not (l_e_st_eq_landau_n_rt_is x0 y0))). Time Defined. (* constant 1417 *) Definition l_e_st_eq_landau_n_rt_some : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)) => l_some l_e_st_eq_landau_n_rt_rat p). Time Defined. (* constant 1418 *) Definition l_e_st_eq_landau_n_rt_all : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)) => l_all l_e_st_eq_landau_n_rt_rat p). Time Defined. (* constant 1419 *) Definition l_e_st_eq_landau_n_rt_one : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)) => l_e_one l_e_st_eq_landau_n_rt_rat p). Time Defined. (* constant 1420 *) Definition l_e_st_eq_landau_n_rt_in : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_esti l_e_st_eq_landau_n_rt_rat x0 s)). Time Defined. (* constant 1421 *) Definition l_e_st_eq_landau_n_rt_ratof : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_rt_rat). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_ectelt l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq x). Time Defined. (* constant 1422 *) Definition l_e_st_eq_landau_n_rt_class : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_set l_e_st_eq_landau_n_frac). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_ecect l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq x0). Time Defined. (* constant 1423 *) Definition l_e_st_eq_landau_n_rt_inclass : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_ratof x))). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_4_th5 l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq x). Time Defined. (* constant 1424 *) Definition l_e_st_eq_landau_n_rt_lemmaeq1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (e:l_e_st_eq_landau_n_eq x y), l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class x0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (e:l_e_st_eq_landau_n_eq x y) => l_e_st_eq_4_th8 l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq x0 x xix0 y e))))). Time Defined. (* constant 1425 *) Definition l_e_st_eq_landau_n_rt_ratapp1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:Prop), (forall (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), a))), a))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:Prop) => (fun (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), a))) => l_e_st_eq_4_th3 l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq x0 a a1))). Time Defined. (* constant 1426 *) Definition l_e_st_eq_landau_n_rt_ii5_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:Prop), (forall (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), a))))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), a)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:Prop) => (fun (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), a))))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_ratapp1 y0 a (fun (y:l_e_st_eq_landau_n_frac) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => a1 x y xix0 yi)))))))). Time Defined. (* constant 1427 *) Definition l_e_st_eq_landau_n_rt_ratapp2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:Prop), (forall (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), a))))), a)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:Prop) => (fun (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), a))))) => l_e_st_eq_landau_n_rt_ratapp1 x0 a (fun (x:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_ii5_t1 x0 y0 a a1 x xi)))))). Time Defined. (* constant 1428 *) Definition l_e_st_eq_landau_n_rt_ii5_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (a:Prop), (forall (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), a))))))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), a))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (a:Prop) => (fun (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), a))))))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_ratapp2 y0 z0 a (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => a1 x y z xix0 yi zi))))))))))). Time Defined. (* constant 1429 *) Definition l_e_st_eq_landau_n_rt_ratapp3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (a:Prop), (forall (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), a))))))), a))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (a:Prop) => (fun (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), a))))))) => l_e_st_eq_landau_n_rt_ratapp1 x0 a (fun (x:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_ii5_t2 x0 y0 z0 a a1 x xi))))))). Time Defined. (* constant 1430 *) Definition l_e_st_eq_landau_n_rt_ii5_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (a:Prop), (forall (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), a))))))))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), a)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (a:Prop) => (fun (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), a))))))))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_ratapp3 y0 z0 u0 a (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => a1 x y z u xix0 yi zi ui)))))))))))))). Time Defined. (* constant 1431 *) Definition l_e_st_eq_landau_n_rt_ratapp4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (a:Prop), (forall (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), a))))))))), a)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (a:Prop) => (fun (a1:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), a))))))))) => l_e_st_eq_landau_n_rt_ratapp1 x0 a (fun (x:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_ii5_t3 x0 y0 z0 u0 a a1 x xi)))))))). Time Defined. (* constant 1432 *) Definition l_e_st_eq_landau_n_rt_isi : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (e:l_e_st_eq_landau_n_eq x1 y1), l_e_st_eq_landau_n_rt_is x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (e:l_e_st_eq_landau_n_eq x1 y1) => l_e_st_eq_5_th3 l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq x0 y0 x1 x1ix0 y1 y1iy0 e))))))). Time Defined. (* constant 1433 *) Definition l_e_st_eq_landau_n_rt_ise : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_eq x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_5_th5 l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq x0 y0 x1 x1ix0 y1 y1iy0 i))))))). Time Defined. (* constant 1434 *) Definition l_e_st_eq_landau_n_rt_nisi : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (n:l_not (l_e_st_eq_landau_n_eq x1 y1)), l_e_st_eq_landau_n_rt_nis x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (n:l_not (l_e_st_eq_landau_n_eq x1 y1)) => l_imp_th3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_eq x1 y1) n (fun (t:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ise x0 y0 x1 y1 x1ix0 y1iy0 t)))))))). Time Defined. (* constant 1435 *) Definition l_e_st_eq_landau_n_rt_nise : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (n:l_e_st_eq_landau_n_rt_nis x0 y0), l_not (l_e_st_eq_landau_n_eq x1 y1)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (n:l_e_st_eq_landau_n_rt_nis x0 y0) => l_imp_th3 (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_rt_is x0 y0) n (fun (t:l_e_st_eq_landau_n_eq x1 y1) => l_e_st_eq_landau_n_rt_isi x0 y0 x1 y1 x1ix0 y1iy0 t)))))))). Time Defined. (* constant 1436 *) Definition l_e_st_eq_landau_n_rt_fixf : (forall (alpha:Type), (forall (f:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), alpha))), Prop)). exact (fun (alpha:Type) => (fun (f:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), alpha))) => l_e_st_eq_fixfu2 l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq alpha f)). Time Defined. (* constant 1437 *) Definition l_e_st_eq_landau_n_rt_indrat : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (alpha:Type), (forall (f:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), alpha))), (forall (ff:l_e_st_eq_landau_n_rt_fixf alpha f), alpha))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (alpha:Type) => (fun (f:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), alpha))) => (fun (ff:l_e_st_eq_landau_n_rt_fixf alpha f) => l_e_st_eq_indeq2 l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq alpha f ff x0 y0))))). Time Defined. (* constant 1438 *) Definition l_e_st_eq_landau_n_rt_isindrat : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (alpha:Type), (forall (f:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), alpha))), (forall (ff:l_e_st_eq_landau_n_rt_fixf alpha f), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_is alpha (f x y) (l_e_st_eq_landau_n_rt_indrat x0 y0 alpha f ff)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (alpha:Type) => (fun (f:(forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), alpha))) => (fun (ff:l_e_st_eq_landau_n_rt_fixf alpha f) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_11_th1 l_e_st_eq_landau_n_frac l_e_st_eq_landau_n_rt_eq l_e_st_eq_landau_n_rt_refeq l_e_st_eq_landau_n_rt_symeq l_e_st_eq_landau_n_rt_treq alpha f ff x0 y0 x xix0 y yiy0))))))))). Time Defined. (* constant 1439 *) Definition l_e_st_eq_landau_n_rt_satz78 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is x0 x0). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_refis l_e_st_eq_landau_n_rt_rat x0). Time Defined. (* constant 1440 *) Definition l_e_st_eq_landau_n_rt_satz79 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_symis l_e_st_eq_landau_n_rt_rat x0 y0 i))). Time Defined. (* constant 1441 *) Definition l_e_st_eq_landau_n_rt_satz80 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is y0 z0), l_e_st_eq_landau_n_rt_is x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is y0 z0) => l_e_tris l_e_st_eq_landau_n_rt_rat x0 y0 z0 i j))))). Time Defined. (* constant 1442 *) Definition l_e_st_eq_landau_n_rt_more : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_some l_e_st_eq_landau_n_frac (fun (x:l_e_st_eq_landau_n_frac) => l_some l_e_st_eq_landau_n_frac (fun (y:l_e_st_eq_landau_n_frac) => l_and3 (l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_moref x y))))). Time Defined. (* constant 1443 *) Definition l_e_st_eq_landau_n_rt_ii5_propm : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), Prop)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => l_and3 (l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_moref x1 y1))))). Time Defined. (* constant 1444 *) Definition l_e_st_eq_landau_n_rt_ii5_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)), (forall (u:l_e_st_eq_landau_n_frac), (forall (p:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u), l_e_st_eq_landau_n_rt_inf t (l_e_st_eq_landau_n_rt_class x0)))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (p:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u) => l_and3e1 (l_e_st_eq_landau_n_rt_inf t (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_moref t u) p))))))))))). Time Defined. (* constant 1445 *) Definition l_e_st_eq_landau_n_rt_ii5_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)), (forall (u:l_e_st_eq_landau_n_frac), (forall (p:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u), l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class y0)))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (p:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u) => l_and3e2 (l_e_st_eq_landau_n_rt_inf t (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_moref t u) p))))))))))). Time Defined. (* constant 1446 *) Definition l_e_st_eq_landau_n_rt_ii5_t6 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)), (forall (u:l_e_st_eq_landau_n_frac), (forall (p:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u), l_e_st_eq_landau_n_moref t u))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (p:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u) => l_and3e3 (l_e_st_eq_landau_n_rt_inf t (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_moref t u) p))))))))))). Time Defined. (* constant 1447 *) Definition l_e_st_eq_landau_n_rt_ii5_t7 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)), (forall (u:l_e_st_eq_landau_n_frac), (forall (p:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u), l_e_st_eq_landau_n_moref x y))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (p:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u) => l_e_st_eq_landau_n_satz44 t u x y (l_e_st_eq_landau_n_rt_ii5_t6 x0 y0 m x y xix0 yiy0 t s u p) (l_e_st_eq_landau_n_rt_ise x0 x0 t x (l_e_st_eq_landau_n_rt_ii5_t4 x0 y0 m x y xix0 yiy0 t s u p) xix0 (l_e_refis l_e_st_eq_landau_n_rt_rat x0)) (l_e_st_eq_landau_n_rt_ise y0 y0 u y (l_e_st_eq_landau_n_rt_ii5_t5 x0 y0 m x y xix0 yiy0 t s u p) yiy0 (l_e_refis l_e_st_eq_landau_n_rt_rat y0))))))))))))). Time Defined. (* constant 1448 *) Definition l_e_st_eq_landau_n_rt_ii5_t8 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)), l_e_st_eq_landau_n_moref x y))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)) => l_someapp l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u) s (l_e_st_eq_landau_n_moref x y) (fun (u:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u) => l_e_st_eq_landau_n_rt_ii5_t7 x0 y0 m x y xix0 yiy0 t s u v))))))))))). Time Defined. (* constant 1449 *) Definition l_e_st_eq_landau_n_rt_also18 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_moref x y))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_someapp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)) m (l_e_st_eq_landau_n_moref x y) (fun (t:l_e_st_eq_landau_n_frac) => (fun (v:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 t u)) => l_e_st_eq_landau_n_rt_ii5_t8 x0 y0 m x y xix0 yiy0 t v))))))))). Time Defined. (* constant 1450 *) Definition l_e_st_eq_landau_n_rt_ii5_t9 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (m:l_e_st_eq_landau_n_moref x1 y1), l_e_st_eq_landau_n_rt_ii5_propm x0 y0 x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (m:l_e_st_eq_landau_n_moref x1 y1) => l_and3i (l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_moref x1 y1) x1ix0 y1iy0 m))))))). Time Defined. (* constant 1451 *) Definition l_e_st_eq_landau_n_rt_ii5_t10 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (m:l_e_st_eq_landau_n_moref x1 y1), l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 x1 t)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (m:l_e_st_eq_landau_n_moref x1 y1) => l_somei l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 x1 t) y1 (l_e_st_eq_landau_n_rt_ii5_t9 x0 y0 x1 y1 x1ix0 y1iy0 m)))))))). Time Defined. (* constant 1452 *) Definition l_e_st_eq_landau_n_rt_morei : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (m:l_e_st_eq_landau_n_moref x1 y1), l_e_st_eq_landau_n_rt_more x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (m:l_e_st_eq_landau_n_moref x1 y1) => l_somei l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propm x0 y0 u t)) x1 (l_e_st_eq_landau_n_rt_ii5_t10 x0 y0 x1 y1 x1ix0 y1iy0 m)))))))). Time Defined. (* constant 1453 *) Definition l_e_st_eq_landau_n_rt_moree : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_moref x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_also18 x0 y0 m x1 y1 x1ix0 y1iy0))))))). Time Defined. (* constant 1454 *) Definition l_e_st_eq_landau_n_rt_ismore1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 z0), l_e_st_eq_landau_n_rt_more y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 z0) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t z0) x0 y0 m i))))). Time Defined. (* constant 1455 *) Definition l_e_st_eq_landau_n_rt_ismore2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (m:l_e_st_eq_landau_n_rt_more z0 x0), l_e_st_eq_landau_n_rt_more z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (m:l_e_st_eq_landau_n_rt_more z0 x0) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more z0 t) x0 y0 m i))))). Time Defined. (* constant 1456 *) Definition l_e_st_eq_landau_n_rt_ismore12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 u0), (forall (m:l_e_st_eq_landau_n_rt_more x0 z0), l_e_st_eq_landau_n_rt_more y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 u0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 z0) => l_e_st_eq_landau_n_rt_ismore2 z0 u0 y0 j (l_e_st_eq_landau_n_rt_ismore1 x0 y0 z0 i m)))))))). Time Defined. (* constant 1457 *) Definition l_e_st_eq_landau_n_rt_less : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_some l_e_st_eq_landau_n_frac (fun (x:l_e_st_eq_landau_n_frac) => l_some l_e_st_eq_landau_n_frac (fun (y:l_e_st_eq_landau_n_frac) => l_and3 (l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_lessf x y))))). Time Defined. (* constant 1458 *) Definition l_e_st_eq_landau_n_rt_ii5_propl : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), Prop)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => l_and3 (l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_lessf x1 y1))))). Time Defined. (* constant 1459 *) Definition l_e_st_eq_landau_n_rt_ii5_t11 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)), (forall (u:l_e_st_eq_landau_n_frac), (forall (p:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u), l_e_st_eq_landau_n_rt_inf t (l_e_st_eq_landau_n_rt_class x0)))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (p:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u) => l_and3e1 (l_e_st_eq_landau_n_rt_inf t (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_lessf t u) p))))))))))). Time Defined. (* constant 1460 *) Definition l_e_st_eq_landau_n_rt_ii5_t12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)), (forall (u:l_e_st_eq_landau_n_frac), (forall (p:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u), l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class y0)))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (p:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u) => l_and3e2 (l_e_st_eq_landau_n_rt_inf t (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_lessf t u) p))))))))))). Time Defined. (* constant 1461 *) Definition l_e_st_eq_landau_n_rt_ii5_t13 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)), (forall (u:l_e_st_eq_landau_n_frac), (forall (p:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u), l_e_st_eq_landau_n_lessf t u))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (p:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u) => l_and3e3 (l_e_st_eq_landau_n_rt_inf t (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_lessf t u) p))))))))))). Time Defined. (* constant 1462 *) Definition l_e_st_eq_landau_n_rt_ii5_t14 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)), (forall (u:l_e_st_eq_landau_n_frac), (forall (p:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u), l_e_st_eq_landau_n_lessf x y))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (p:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u) => l_e_st_eq_landau_n_satz45 t u x y (l_e_st_eq_landau_n_rt_ii5_t13 x0 y0 l x y xix0 yiy0 t s u p) (l_e_st_eq_landau_n_rt_ise x0 x0 t x (l_e_st_eq_landau_n_rt_ii5_t11 x0 y0 l x y xix0 yiy0 t s u p) xix0 (l_e_refis l_e_st_eq_landau_n_rt_rat x0)) (l_e_st_eq_landau_n_rt_ise y0 y0 u y (l_e_st_eq_landau_n_rt_ii5_t12 x0 y0 l x y xix0 yiy0 t s u p) yiy0 (l_e_refis l_e_st_eq_landau_n_rt_rat y0))))))))))))). Time Defined. (* constant 1463 *) Definition l_e_st_eq_landau_n_rt_ii5_t15 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (t:l_e_st_eq_landau_n_frac), (forall (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)), l_e_st_eq_landau_n_lessf x y))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (t:l_e_st_eq_landau_n_frac) => (fun (s:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)) => l_someapp l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u) s (l_e_st_eq_landau_n_lessf x y) (fun (u:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u) => l_e_st_eq_landau_n_rt_ii5_t14 x0 y0 l x y xix0 yiy0 t s u v))))))))))). Time Defined. (* constant 1464 *) Definition l_e_st_eq_landau_n_rt_also19 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_lessf x y))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_someapp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)) l (l_e_st_eq_landau_n_lessf x y) (fun (t:l_e_st_eq_landau_n_frac) => (fun (v:l_some l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 t u)) => l_e_st_eq_landau_n_rt_ii5_t15 x0 y0 l x y xix0 yiy0 t v))))))))). Time Defined. (* constant 1465 *) Definition l_e_st_eq_landau_n_rt_ii5_t16 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (l:l_e_st_eq_landau_n_lessf x1 y1), l_e_st_eq_landau_n_rt_ii5_propl x0 y0 x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (l:l_e_st_eq_landau_n_lessf x1 y1) => l_and3i (l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) (l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) (l_e_st_eq_landau_n_lessf x1 y1) x1ix0 y1iy0 l))))))). Time Defined. (* constant 1466 *) Definition l_e_st_eq_landau_n_rt_ii5_t17 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (l:l_e_st_eq_landau_n_lessf x1 y1), l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 x1 t)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (l:l_e_st_eq_landau_n_lessf x1 y1) => l_somei l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 x1 t) y1 (l_e_st_eq_landau_n_rt_ii5_t16 x0 y0 x1 y1 x1ix0 y1iy0 l)))))))). Time Defined. (* constant 1467 *) Definition l_e_st_eq_landau_n_rt_lessi : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (l:l_e_st_eq_landau_n_lessf x1 y1), l_e_st_eq_landau_n_rt_less x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (l:l_e_st_eq_landau_n_lessf x1 y1) => l_somei l_e_st_eq_landau_n_frac (fun (u:l_e_st_eq_landau_n_frac) => l_some l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ii5_propl x0 y0 u t)) x1 (l_e_st_eq_landau_n_rt_ii5_t17 x0 y0 x1 y1 x1ix0 y1iy0 l)))))))). Time Defined. (* constant 1468 *) Definition l_e_st_eq_landau_n_rt_lesse : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_lessf x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_also19 x0 y0 l x1 y1 x1ix0 y1iy0))))))). Time Defined. (* constant 1469 *) Definition l_e_st_eq_landau_n_rt_isless1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (l:l_e_st_eq_landau_n_rt_less x0 z0), l_e_st_eq_landau_n_rt_less y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (l:l_e_st_eq_landau_n_rt_less x0 z0) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less t z0) x0 y0 l i))))). Time Defined. (* constant 1470 *) Definition l_e_st_eq_landau_n_rt_isless2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (l:l_e_st_eq_landau_n_rt_less z0 x0), l_e_st_eq_landau_n_rt_less z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (l:l_e_st_eq_landau_n_rt_less z0 x0) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less z0 t) x0 y0 l i))))). Time Defined. (* constant 1471 *) Definition l_e_st_eq_landau_n_rt_isless12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 u0), (forall (l:l_e_st_eq_landau_n_rt_less x0 z0), l_e_st_eq_landau_n_rt_less y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 u0) => (fun (l:l_e_st_eq_landau_n_rt_less x0 z0) => l_e_st_eq_landau_n_rt_isless2 z0 u0 y0 j (l_e_st_eq_landau_n_rt_isless1 x0 y0 z0 i l)))))))). Time Defined. (* constant 1472 *) Definition l_e_st_eq_landau_n_rt_581_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_or3 (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_moref x1 y1) (l_e_st_eq_landau_n_lessf x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_satz41a x1 y1)))))). Time Defined. (* constant 1473 *) Definition l_e_st_eq_landau_n_rt_581_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (e:l_e_st_eq_landau_n_eq x1 y1), l_or3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (e:l_e_st_eq_landau_n_eq x1 y1) => l_or3i1 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_isi x0 y0 x1 y1 x1ix0 y1iy0 e)))))))). Time Defined. (* constant 1474 *) Definition l_e_st_eq_landau_n_rt_581_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (m:l_e_st_eq_landau_n_moref x1 y1), l_or3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (m:l_e_st_eq_landau_n_moref x1 y1) => l_or3i2 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_morei x0 y0 x1 y1 x1ix0 y1iy0 m)))))))). Time Defined. (* constant 1475 *) Definition l_e_st_eq_landau_n_rt_581_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (l:l_e_st_eq_landau_n_lessf x1 y1), l_or3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (l:l_e_st_eq_landau_n_lessf x1 y1) => l_or3i3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_lessi x0 y0 x1 y1 x1ix0 y1iy0 l)))))))). Time Defined. (* constant 1476 *) Definition l_e_st_eq_landau_n_rt_581_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_or3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_or3app (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_moref x1 y1) (l_e_st_eq_landau_n_lessf x1 y1) (l_or3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0)) (l_e_st_eq_landau_n_rt_581_t1 x0 y0 x1 y1 x1ix0 y1iy0) (fun (t:l_e_st_eq_landau_n_eq x1 y1) => l_e_st_eq_landau_n_rt_581_t2 x0 y0 x1 y1 x1ix0 y1iy0 t) (fun (t:l_e_st_eq_landau_n_moref x1 y1) => l_e_st_eq_landau_n_rt_581_t3 x0 y0 x1 y1 x1ix0 y1iy0 t) (fun (t:l_e_st_eq_landau_n_lessf x1 y1) => l_e_st_eq_landau_n_rt_581_t4 x0 y0 x1 y1 x1ix0 y1iy0 t))))))). Time Defined. (* constant 1477 *) Definition l_e_st_eq_landau_n_rt_581_t6 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_ec3 (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_moref x1 y1) (l_e_st_eq_landau_n_lessf x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_satz41b x1 y1)))))). Time Defined. (* constant 1478 *) Definition l_e_st_eq_landau_n_rt_581_t7 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_not (l_e_st_eq_landau_n_rt_more x0 y0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_imp_th3 (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_moref x1 y1) (l_ec3e12 (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_moref x1 y1) (l_e_st_eq_landau_n_lessf x1 y1) (l_e_st_eq_landau_n_rt_581_t6 x0 y0 x1 y1 x1ix0 y1iy0) (l_e_st_eq_landau_n_rt_ise x0 y0 x1 y1 x1ix0 y1iy0 i)) (fun (t:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_moree x0 y0 x1 y1 x1ix0 y1iy0 t)))))))). Time Defined. (* constant 1479 *) Definition l_e_st_eq_landau_n_rt_581_t8 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_not (l_e_st_eq_landau_n_rt_less x0 y0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_imp_th3 (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_lessf x1 y1) (l_ec3e23 (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_moref x1 y1) (l_e_st_eq_landau_n_lessf x1 y1) (l_e_st_eq_landau_n_rt_581_t6 x0 y0 x1 y1 x1ix0 y1iy0) (l_e_st_eq_landau_n_rt_moree x0 y0 x1 y1 x1ix0 y1iy0 m)) (fun (t:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_lesse x0 y0 x1 y1 x1ix0 y1iy0 t)))))))). Time Defined. (* constant 1480 *) Definition l_e_st_eq_landau_n_rt_581_t9 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_nis x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_imp_th3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_eq x1 y1) (l_ec3e31 (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_moref x1 y1) (l_e_st_eq_landau_n_lessf x1 y1) (l_e_st_eq_landau_n_rt_581_t6 x0 y0 x1 y1 x1ix0 y1iy0) (l_e_st_eq_landau_n_rt_lesse x0 y0 x1 y1 x1ix0 y1iy0 l)) (fun (t:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ise x0 y0 x1 y1 x1ix0 y1iy0 t)))))))). Time Defined. (* constant 1481 *) Definition l_e_st_eq_landau_n_rt_581_t10 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_ec3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_ec3_th6 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_ec_th1 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (fun (t:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_581_t7 x0 y0 x1 y1 x1ix0 y1iy0 t)) (l_ec_th1 (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (fun (t:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_581_t8 x0 y0 x1 y1 x1ix0 y1iy0 t)) (l_ec_th1 (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) (fun (t:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_581_t9 x0 y0 x1 y1 x1ix0 y1iy0 t)))))))). Time Defined. (* constant 1482 *) Definition l_e_st_eq_landau_n_rt_581_t11 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_orec3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_orec3i (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_581_t5 x0 y0 x1 y1 x1ix0 y1iy0) (l_e_st_eq_landau_n_rt_581_t10 x0 y0 x1 y1 x1ix0 y1iy0))))))). Time Defined. (* constant 1483 *) Definition l_e_st_eq_landau_n_rt_satz81 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_orec3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_orec3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_581_t11 x0 y0 x y xi yi)))))). Time Defined. (* constant 1484 *) Definition l_e_st_eq_landau_n_rt_satz81a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_or3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_orec3e1 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81 x0 y0))). Time Defined. (* constant 1485 *) Definition l_e_st_eq_landau_n_rt_satz81b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_ec3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_orec3e2 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81 x0 y0))). Time Defined. (* constant 1486 *) Definition l_e_st_eq_landau_n_rt_582_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_less y0 x0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_lessi y0 x0 y x yiy0 xix0 (l_e_st_eq_landau_n_satz42 x y (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m))))))))). Time Defined. (* constant 1487 *) Definition l_e_st_eq_landau_n_rt_satz82 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_less y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_less y0 x0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_582_t1 x0 y0 m x y xi yi))))))). Time Defined. (* constant 1488 *) Definition l_e_st_eq_landau_n_rt_583_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_more y0 x0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_morei y0 x0 y x yiy0 xix0 (l_e_st_eq_landau_n_satz43 x y (l_e_st_eq_landau_n_rt_lesse x0 y0 x y xix0 yiy0 l))))))))). Time Defined. (* constant 1489 *) Definition l_e_st_eq_landau_n_rt_satz83 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_more y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_more y0 x0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_583_t1 x0 y0 l x y xi yi))))))). Time Defined. (* constant 1490 *) Definition l_e_st_eq_landau_n_rt_moreis : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_or (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0))). Time Defined. (* constant 1491 *) Definition l_e_st_eq_landau_n_rt_moreisi1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_moreis x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_ori1 (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) m))). Time Defined. (* constant 1492 *) Definition l_e_st_eq_landau_n_rt_moreisi2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_moreis x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_ori2 (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) i))). Time Defined. (* constant 1493 *) Definition l_e_st_eq_landau_n_rt_moreisi : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (m:l_e_st_eq_landau_n_moreq x1 y1), l_e_st_eq_landau_n_rt_moreis x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (m:l_e_st_eq_landau_n_moreq x1 y1) => l_orapp (l_e_st_eq_landau_n_moref x1 y1) (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_rt_moreis x0 y0) m (fun (t:l_e_st_eq_landau_n_moref x1 y1) => l_e_st_eq_landau_n_rt_moreisi1 x0 y0 (l_e_st_eq_landau_n_rt_morei x0 y0 x1 y1 x1ix0 y1iy0 t)) (fun (t:l_e_st_eq_landau_n_eq x1 y1) => l_e_st_eq_landau_n_rt_moreisi2 x0 y0 (l_e_st_eq_landau_n_rt_isi x0 y0 x1 y1 x1ix0 y1iy0 t))))))))). Time Defined. (* constant 1494 *) Definition l_e_st_eq_landau_n_rt_moreise : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), l_e_st_eq_landau_n_moreq x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => l_orapp (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_moreq x1 y1) m (fun (t:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_moreqi1 x1 y1 (l_e_st_eq_landau_n_rt_moree x0 y0 x1 y1 x1ix0 y1iy0 t)) (fun (t:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_moreqi2 x1 y1 (l_e_st_eq_landau_n_rt_ise x0 y0 x1 y1 x1ix0 y1iy0 t))))))))). Time Defined. (* constant 1495 *) Definition l_e_st_eq_landau_n_rt_ismoreis1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 z0), l_e_st_eq_landau_n_rt_moreis y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 z0) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_moreis t z0) x0 y0 m i))))). Time Defined. (* constant 1496 *) Definition l_e_st_eq_landau_n_rt_ismoreis2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (m:l_e_st_eq_landau_n_rt_moreis z0 x0), l_e_st_eq_landau_n_rt_moreis z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (m:l_e_st_eq_landau_n_rt_moreis z0 x0) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_moreis z0 t) x0 y0 m i))))). Time Defined. (* constant 1497 *) Definition l_e_st_eq_landau_n_rt_ismoreis12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 u0), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 z0), l_e_st_eq_landau_n_rt_moreis y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 u0) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 z0) => l_e_st_eq_landau_n_rt_ismoreis2 z0 u0 y0 j (l_e_st_eq_landau_n_rt_ismoreis1 x0 y0 z0 i m)))))))). Time Defined. (* constant 1498 *) Definition l_e_st_eq_landau_n_rt_lessis : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_or (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0))). Time Defined. (* constant 1499 *) Definition l_e_st_eq_landau_n_rt_lessisi1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_lessis x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_ori1 (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) l))). Time Defined. (* constant 1500 *) Definition l_e_st_eq_landau_n_rt_lessisi2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_lessis x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_ori2 (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) i))). Time Defined. (* constant 1501 *) Definition l_e_st_eq_landau_n_rt_lessisi : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (l:l_e_st_eq_landau_n_lesseq x1 y1), l_e_st_eq_landau_n_rt_lessis x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (l:l_e_st_eq_landau_n_lesseq x1 y1) => l_orapp (l_e_st_eq_landau_n_lessf x1 y1) (l_e_st_eq_landau_n_eq x1 y1) (l_e_st_eq_landau_n_rt_lessis x0 y0) l (fun (t:l_e_st_eq_landau_n_lessf x1 y1) => l_e_st_eq_landau_n_rt_lessisi1 x0 y0 (l_e_st_eq_landau_n_rt_lessi x0 y0 x1 y1 x1ix0 y1iy0 t)) (fun (t:l_e_st_eq_landau_n_eq x1 y1) => l_e_st_eq_landau_n_rt_lessisi2 x0 y0 (l_e_st_eq_landau_n_rt_isi x0 y0 x1 y1 x1ix0 y1iy0 t))))))))). Time Defined. (* constant 1502 *) Definition l_e_st_eq_landau_n_rt_lessise : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), l_e_st_eq_landau_n_lesseq x1 y1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => l_orapp (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_lesseq x1 y1) l (fun (t:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_lesseqi1 x1 y1 (l_e_st_eq_landau_n_rt_lesse x0 y0 x1 y1 x1ix0 y1iy0 t)) (fun (t:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_lesseqi2 x1 y1 (l_e_st_eq_landau_n_rt_ise x0 y0 x1 y1 x1ix0 y1iy0 t))))))))). Time Defined. (* constant 1503 *) Definition l_e_st_eq_landau_n_rt_islessis1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 z0), l_e_st_eq_landau_n_rt_lessis y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 z0) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lessis t z0) x0 y0 l i))))). Time Defined. (* constant 1504 *) Definition l_e_st_eq_landau_n_rt_islessis2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (l:l_e_st_eq_landau_n_rt_lessis z0 x0), l_e_st_eq_landau_n_rt_lessis z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (l:l_e_st_eq_landau_n_rt_lessis z0 x0) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lessis z0 t) x0 y0 l i))))). Time Defined. (* constant 1505 *) Definition l_e_st_eq_landau_n_rt_islessis12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 u0), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 z0), l_e_st_eq_landau_n_rt_lessis y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 u0) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 z0) => l_e_st_eq_landau_n_rt_islessis2 z0 u0 y0 j (l_e_st_eq_landau_n_rt_islessis1 x0 y0 z0 i l)))))))). Time Defined. (* constant 1506 *) Definition l_e_st_eq_landau_n_rt_satz81c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), l_not (l_e_st_eq_landau_n_rt_less x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => l_ec3_th7 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) (l_comor (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) m)))). Time Defined. (* constant 1507 *) Definition l_e_st_eq_landau_n_rt_satz81d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), l_not (l_e_st_eq_landau_n_rt_more x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => l_ec3_th9 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) l))). Time Defined. (* constant 1508 *) Definition l_e_st_eq_landau_n_rt_satz81e : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_more x0 y0)), l_e_st_eq_landau_n_rt_lessis x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_more x0 y0)) => l_or3_th2 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81a x0 y0) n))). Time Defined. (* constant 1509 *) Definition l_e_st_eq_landau_n_rt_satz81f : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_less x0 y0)), l_e_st_eq_landau_n_rt_moreis x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_less x0 y0)) => l_comor (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_or3_th3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81a x0 y0) n)))). Time Defined. (* constant 1510 *) Definition l_e_st_eq_landau_n_rt_satz81g : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_not (l_e_st_eq_landau_n_rt_lessis x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_or_th3 (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) (l_ec3e23 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) m) (l_ec3e21 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) m)))). Time Defined. (* constant 1511 *) Definition l_e_st_eq_landau_n_rt_satz81h : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_not (l_e_st_eq_landau_n_rt_moreis x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_or_th3 (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) (l_ec3e32 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) l) (l_ec3e31 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) l)))). Time Defined. (* constant 1512 *) Definition l_e_st_eq_landau_n_rt_satz81j : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 y0)), l_e_st_eq_landau_n_rt_less x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 y0)) => l_or3e3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81a x0 y0) (l_or_th5 (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) n) (l_or_th4 (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) n)))). Time Defined. (* constant 1513 *) Definition l_e_st_eq_landau_n_rt_satz81k : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_lessis x0 y0)), l_e_st_eq_landau_n_rt_more x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_lessis x0 y0)) => l_or3e2 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81a x0 y0) (l_or_th4 (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) n) (l_or_th5 (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) n)))). Time Defined. (* constant 1514 *) Definition l_e_st_eq_landau_n_rt_584_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_lessis y0 x0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_lessisi y0 x0 y x yiy0 xix0 (l_e_st_eq_landau_n_satz48 x y (l_e_st_eq_landau_n_rt_moreise x0 y0 x y xix0 yiy0 m))))))))). Time Defined. (* constant 1515 *) Definition l_e_st_eq_landau_n_rt_satz84 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), l_e_st_eq_landau_n_rt_lessis y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_lessis y0 x0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_584_t1 x0 y0 m x y xi yi))))))). Time Defined. (* constant 1516 *) Definition l_e_st_eq_landau_n_rt_585_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_moreis y0 x0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_moreisi y0 x0 y x yiy0 xix0 (l_e_st_eq_landau_n_satz49 x y (l_e_st_eq_landau_n_rt_lessise x0 y0 x y xix0 yiy0 l))))))))). Time Defined. (* constant 1517 *) Definition l_e_st_eq_landau_n_rt_satz85 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), l_e_st_eq_landau_n_rt_moreis y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_moreis y0 x0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_585_t1 x0 y0 l x y xi yi))))))). Time Defined. (* constant 1518 *) Definition l_e_st_eq_landau_n_rt_586_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less y0 z0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_less x0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less y0 z0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_lessi x0 z0 x z xix0 ziz0 (l_e_st_eq_landau_n_satz50 x y z (l_e_st_eq_landau_n_rt_lesse x0 y0 x y xix0 yiy0 l) (l_e_st_eq_landau_n_rt_lesse y0 z0 y z yiy0 ziz0 k))))))))))))). Time Defined. (* constant 1519 *) Definition l_e_st_eq_landau_n_rt_satz86 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less y0 z0), l_e_st_eq_landau_n_rt_less x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less y0 z0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_less x0 z0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_586_t1 x0 y0 z0 l k x y z xi yi zi))))))))))). Time Defined. (* constant 1520 *) Definition l_e_st_eq_landau_n_rt_trless : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less y0 z0), l_e_st_eq_landau_n_rt_less x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less y0 z0) => l_e_st_eq_landau_n_rt_satz86 x0 y0 z0 l k))))). Time Defined. (* constant 1521 *) Definition l_e_st_eq_landau_n_rt_trmore : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more y0 z0), l_e_st_eq_landau_n_rt_more x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more y0 z0) => l_e_st_eq_landau_n_rt_satz83 z0 x0 (l_e_st_eq_landau_n_rt_satz86 z0 y0 x0 (l_e_st_eq_landau_n_rt_satz82 y0 z0 n) (l_e_st_eq_landau_n_rt_satz82 x0 y0 m))))))). Time Defined. (* constant 1522 *) Definition l_e_st_eq_landau_n_rt_587_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less y0 z0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_less x0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less y0 z0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_lessi x0 z0 x z xix0 ziz0 (l_e_st_eq_landau_n_satz51a x y z (l_e_st_eq_landau_n_rt_lessise x0 y0 x y xix0 yiy0 l) (l_e_st_eq_landau_n_rt_lesse y0 z0 y z yiy0 ziz0 k))))))))))))). Time Defined. (* constant 1523 *) Definition l_e_st_eq_landau_n_rt_satz87a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less y0 z0), l_e_st_eq_landau_n_rt_less x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less y0 z0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_less x0 z0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_587_t1 x0 y0 z0 l k x y z xi yi zi))))))))))). Time Defined. (* constant 1524 *) Definition l_e_st_eq_landau_n_rt_587_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis y0 z0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_less x0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis y0 z0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_lessi x0 z0 x z xix0 ziz0 (l_e_st_eq_landau_n_satz51b x y z (l_e_st_eq_landau_n_rt_lesse x0 y0 x y xix0 yiy0 l) (l_e_st_eq_landau_n_rt_lessise y0 z0 y z yiy0 ziz0 k))))))))))))). Time Defined. (* constant 1525 *) Definition l_e_st_eq_landau_n_rt_satz87b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis y0 z0), l_e_st_eq_landau_n_rt_less x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis y0 z0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_less x0 z0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_587_t2 x0 y0 z0 l k x y z xi yi zi))))))))))). Time Defined. (* constant 1526 *) Definition l_e_st_eq_landau_n_rt_satz87c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more y0 z0), l_e_st_eq_landau_n_rt_more x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more y0 z0) => l_e_st_eq_landau_n_rt_satz83 z0 x0 (l_e_st_eq_landau_n_rt_satz87b z0 y0 x0 (l_e_st_eq_landau_n_rt_satz82 y0 z0 n) (l_e_st_eq_landau_n_rt_satz84 x0 y0 m))))))). Time Defined. (* constant 1527 *) Definition l_e_st_eq_landau_n_rt_satz87d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis y0 z0), l_e_st_eq_landau_n_rt_more x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis y0 z0) => l_e_st_eq_landau_n_rt_satz83 z0 x0 (l_e_st_eq_landau_n_rt_satz87a z0 y0 x0 (l_e_st_eq_landau_n_rt_satz84 y0 z0 n) (l_e_st_eq_landau_n_rt_satz82 x0 y0 m))))))). Time Defined. (* constant 1528 *) Definition l_e_st_eq_landau_n_rt_588_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis y0 z0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_lessis x0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis y0 z0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_lessisi x0 z0 x z xix0 ziz0 (l_e_st_eq_landau_n_satz52 x y z (l_e_st_eq_landau_n_rt_lessise x0 y0 x y xix0 yiy0 l) (l_e_st_eq_landau_n_rt_lessise y0 z0 y z yiy0 ziz0 k))))))))))))). Time Defined. (* constant 1529 *) Definition l_e_st_eq_landau_n_rt_satz88 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis y0 z0), l_e_st_eq_landau_n_rt_lessis x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis y0 z0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_lessis x0 z0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_588_t1 x0 y0 z0 l k x y z xi yi zi))))))))))). Time Defined. (* constant 1530 *) Definition l_e_st_eq_landau_n_rt_trlessis : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis y0 z0), l_e_st_eq_landau_n_rt_lessis x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis y0 z0) => l_e_st_eq_landau_n_rt_satz88 x0 y0 z0 l k))))). Time Defined. (* constant 1531 *) Definition l_e_st_eq_landau_n_rt_trmoreis : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis y0 z0), l_e_st_eq_landau_n_rt_moreis x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis y0 z0) => l_e_st_eq_landau_n_rt_satz85 z0 x0 (l_e_st_eq_landau_n_rt_satz88 z0 y0 x0 (l_e_st_eq_landau_n_rt_satz84 y0 z0 n) (l_e_st_eq_landau_n_rt_satz84 x0 y0 m))))))). Time Defined. (* constant 1532 *) Definition l_e_st_eq_landau_n_rt_589_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (z:l_e_st_eq_landau_n_frac), (forall (m:l_e_st_eq_landau_n_moref z x), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t x0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (m:l_e_st_eq_landau_n_moref z x) => l_somei l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t x0) (l_e_st_eq_landau_n_rt_ratof z) (l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_ratof z) x0 z x (l_e_st_eq_landau_n_rt_inclass z) xix0 m)))))). Time Defined. (* constant 1533 *) Definition l_e_st_eq_landau_n_rt_589_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t x0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_someapp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_moref t x) (l_e_st_eq_landau_n_satz53 x) (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t x0)) (fun (t:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_moref t x) => l_e_st_eq_landau_n_rt_589_t1 x0 x xix0 t u))))). Time Defined. (* constant 1534 *) Definition l_e_st_eq_landau_n_rt_satz89 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t x0)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp1 x0 (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t x0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_589_t2 x0 x xi))). Time Defined. (* constant 1535 *) Definition l_e_st_eq_landau_n_rt_590_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (z:l_e_st_eq_landau_n_frac), (forall (l:l_e_st_eq_landau_n_lessf z x), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less t x0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (l:l_e_st_eq_landau_n_lessf z x) => l_somei l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less t x0) (l_e_st_eq_landau_n_rt_ratof z) (l_e_st_eq_landau_n_rt_lessi (l_e_st_eq_landau_n_rt_ratof z) x0 z x (l_e_st_eq_landau_n_rt_inclass z) xix0 l)))))). Time Defined. (* constant 1536 *) Definition l_e_st_eq_landau_n_rt_590_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less t x0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_someapp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_lessf t x) (l_e_st_eq_landau_n_satz54 x) (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less t x0)) (fun (t:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_lessf t x) => l_e_st_eq_landau_n_rt_590_t1 x0 x xix0 t u))))). Time Defined. (* constant 1537 *) Definition l_e_st_eq_landau_n_rt_satz90 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less t x0)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp1 x0 (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less t x0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_590_t2 x0 x xi))). Time Defined. (* constant 1538 *) Definition l_e_st_eq_landau_n_rt_591_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (z:l_e_st_eq_landau_n_frac), (forall (a:l_and (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y)), l_e_st_eq_landau_n_rt_less x0 (l_e_st_eq_landau_n_rt_ratof z)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (a:l_and (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y)) => l_e_st_eq_landau_n_rt_lessi x0 (l_e_st_eq_landau_n_rt_ratof z) x z xix0 (l_e_st_eq_landau_n_rt_inclass z) (l_ande1 (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y) a)))))))))). Time Defined. (* constant 1539 *) Definition l_e_st_eq_landau_n_rt_591_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (z:l_e_st_eq_landau_n_frac), (forall (a:l_and (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ratof z) y0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (a:l_and (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y)) => l_e_st_eq_landau_n_rt_lessi (l_e_st_eq_landau_n_rt_ratof z) y0 z y (l_e_st_eq_landau_n_rt_inclass z) yiy0 (l_ande2 (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y) a)))))))))). Time Defined. (* constant 1540 *) Definition l_e_st_eq_landau_n_rt_591_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (z:l_e_st_eq_landau_n_frac), (forall (a:l_and (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y)), l_and (l_e_st_eq_landau_n_rt_less x0 (l_e_st_eq_landau_n_rt_ratof z)) (l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ratof z) y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (a:l_and (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y)) => l_andi (l_e_st_eq_landau_n_rt_less x0 (l_e_st_eq_landau_n_rt_ratof z)) (l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ratof z) y0) (l_e_st_eq_landau_n_rt_591_t1 x0 y0 l x y xix0 yiy0 z a) (l_e_st_eq_landau_n_rt_591_t2 x0 y0 l x y xix0 yiy0 z a)))))))))). Time Defined. (* constant 1541 *) Definition l_e_st_eq_landau_n_rt_591_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (z:l_e_st_eq_landau_n_frac), (forall (a:l_and (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y)), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x0 t) (l_e_st_eq_landau_n_rt_less t y0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (a:l_and (l_e_st_eq_landau_n_lessf x z) (l_e_st_eq_landau_n_lessf z y)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x0 t) (l_e_st_eq_landau_n_rt_less t y0)) (l_e_st_eq_landau_n_rt_ratof z) (l_e_st_eq_landau_n_rt_591_t3 x0 y0 l x y xix0 yiy0 z a)))))))))). Time Defined. (* constant 1542 *) Definition l_e_st_eq_landau_n_rt_591_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x0 t) (l_e_st_eq_landau_n_rt_less t y0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_someapp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_and (l_e_st_eq_landau_n_lessf x t) (l_e_st_eq_landau_n_lessf t y)) (l_e_st_eq_landau_n_satz55 x y (l_e_st_eq_landau_n_rt_lesse x0 y0 x y xix0 yiy0 l)) (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x0 t) (l_e_st_eq_landau_n_rt_less t y0))) (fun (t:l_e_st_eq_landau_n_frac) => (fun (u:l_and (l_e_st_eq_landau_n_lessf x t) (l_e_st_eq_landau_n_lessf t y)) => l_e_st_eq_landau_n_rt_591_t4 x0 y0 l x y xix0 yiy0 t u))))))))). Time Defined. (* constant 1543 *) Definition l_e_st_eq_landau_n_rt_satz91 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x0 t) (l_e_st_eq_landau_n_rt_less t y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x0 t) (l_e_st_eq_landau_n_rt_less t y0))) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_591_t5 x0 y0 l x y xi yi))))))). Time Defined. (* constant 1544 *) Definition l_e_st_eq_landau_n_rt_plusfrt : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_rt_rat)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_pf x y))). Time Defined. (* constant 1545 *) Definition l_e_st_eq_landau_n_rt_ii5_t18 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_plusfrt x z) (l_e_st_eq_landau_n_rt_plusfrt y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_pf x z)) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_pf y u)) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_pf x z)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_pf y u)) (l_e_st_eq_landau_n_satz56 x y z u e f))))))). Time Defined. (* constant 1546 *) Definition l_e_st_eq_landau_n_rt_fplusfrt : l_e_st_eq_landau_n_rt_fixf l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_plusfrt. exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_rt_eq x y) => (fun (w:l_e_st_eq_landau_n_rt_eq z u) => l_e_st_eq_landau_n_rt_ii5_t18 x y z u v w)))))). Time Defined. (* constant 1547 *) Definition l_e_st_eq_landau_n_rt_pl : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rat)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_indrat x0 y0 l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_plusfrt l_e_st_eq_landau_n_rt_fplusfrt)). Time Defined. (* constant 1548 *) Definition l_e_st_eq_landau_n_rt_ii5_t19 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_pf x1 y1)) (l_e_st_eq_landau_n_rt_pl x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_isindrat x0 y0 l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_plusfrt l_e_st_eq_landau_n_rt_fplusfrt x1 y1 x1ix0 y1iy0)))))). Time Defined. (* constant 1549 *) Definition l_e_st_eq_landau_n_rt_picp : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_pf x1 y1) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_pl x0 y0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_pf x1 y1) (l_e_st_eq_landau_n_rt_class t)) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_pf x1 y1)) (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_pf x1 y1)) (l_e_st_eq_landau_n_rt_ii5_t19 x0 y0 x1 y1 x1ix0 y1iy0))))))). Time Defined. (* constant 1550 *) Definition l_e_st_eq_landau_n_rt_ispl1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_isf l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_pl t z0) x0 y0 i)))). Time Defined. (* constant 1551 *) Definition l_e_st_eq_landau_n_rt_ispl2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl z0 x0) (l_e_st_eq_landau_n_rt_pl z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_isf l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_pl z0 t) x0 y0 i)))). Time Defined. (* constant 1552 *) Definition l_e_st_eq_landau_n_rt_ispl12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 u0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 u0) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_rt_ispl1 x0 y0 z0 i) (l_e_st_eq_landau_n_rt_ispl2 z0 u0 y0 j))))))). Time Defined. (* constant 1553 *) Definition l_e_st_eq_landau_n_rt_592_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_pl y0 x0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_pl y0 x0) (l_e_st_eq_landau_n_pf x1 y1) (l_e_st_eq_landau_n_pf y1 x1) (l_e_st_eq_landau_n_rt_picp x0 y0 x1 y1 x1ix0 y1iy0) (l_e_st_eq_landau_n_rt_picp y0 x0 y1 x1 y1iy0 x1ix0) (l_e_st_eq_landau_n_satz58 x1 y1))))))). Time Defined. (* constant 1554 *) Definition l_e_st_eq_landau_n_rt_satz92 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_pl y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_pl y0 x0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_592_t1 x0 y0 x y xi yi)))))). Time Defined. (* constant 1555 *) Definition l_e_st_eq_landau_n_rt_compl : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_pl y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz92 x0 y0)). Time Defined. (* constant 1556 *) Definition l_e_st_eq_landau_n_rt_593_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_picp (l_e_st_eq_landau_n_rt_pl x0 y0) z0 (l_e_st_eq_landau_n_pf x y) z (l_e_st_eq_landau_n_rt_picp x0 y0 x y xix0 yiy0) ziz0))))))))). Time Defined. (* constant 1557 *) Definition l_e_st_eq_landau_n_rt_593_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0)))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_picp x0 (l_e_st_eq_landau_n_rt_pl y0 z0) x (l_e_st_eq_landau_n_pf y z) xix0 (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0)))))))))). Time Defined. (* constant 1558 *) Definition l_e_st_eq_landau_n_rt_593_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_pf x y) z) (l_e_st_eq_landau_n_pf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_rt_593_t1 x0 y0 z0 x y z xix0 yiy0 ziz0) (l_e_st_eq_landau_n_rt_593_t2 x0 y0 z0 x y z xix0 yiy0 ziz0) (l_e_st_eq_landau_n_satz59 x y z)))))))))). Time Defined. (* constant 1559 *) Definition l_e_st_eq_landau_n_rt_satz93 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0))) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_593_t3 x0 y0 z0 x y z xi yi zi))))))))). Time Defined. (* constant 1560 *) Definition l_e_st_eq_landau_n_rt_asspl1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz93 x0 y0 z0))). Time Defined. (* constant 1561 *) Definition l_e_st_eq_landau_n_rt_asspl2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_satz93 x0 y0 z0)))). Time Defined. (* constant 1562 *) Definition l_e_st_eq_landau_n_rt_594_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 y0) x0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_pl x0 y0) x0 (l_e_st_eq_landau_n_pf x1 y1) x1 (l_e_st_eq_landau_n_rt_picp x0 y0 x1 y1 x1ix0 y1iy0) x1ix0 (l_e_st_eq_landau_n_satz60 x1 y1))))))). Time Defined. (* constant 1563 *) Definition l_e_st_eq_landau_n_rt_satz94 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 y0) x0)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 y0) x0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_594_t1 x0 y0 x y xi yi)))))). Time Defined. (* constant 1564 *) Definition l_e_st_eq_landau_n_rt_satz94a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_less x0 (l_e_st_eq_landau_n_rt_pl x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_pl x0 y0) x0 (l_e_st_eq_landau_n_rt_satz94 x0 y0))). Time Defined. (* constant 1565 *) Definition l_e_st_eq_landau_n_rt_595_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0) (l_e_st_eq_landau_n_satz61 x y z (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m)))))))))))). Time Defined. (* constant 1566 *) Definition l_e_st_eq_landau_n_rt_satz95 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_595_t1 x0 y0 z0 m x y z xi yi zi)))))))))). Time Defined. (* constant 1567 *) Definition l_e_st_eq_landau_n_rt_596_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0) (l_e_st_eq_landau_n_satz62a x y z (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m)))))))))))). Time Defined. (* constant 1568 *) Definition l_e_st_eq_landau_n_rt_satz96a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_596_t1 x0 y0 z0 m x y z xi yi zi)))))))))). Time Defined. (* constant 1569 *) Definition l_e_st_eq_landau_n_rt_596_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0) (l_e_st_eq_landau_n_satz62b x y z (l_e_st_eq_landau_n_rt_ise x0 y0 x y xix0 yiy0 i)))))))))))). Time Defined. (* constant 1570 *) Definition l_e_st_eq_landau_n_rt_satz96b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_596_t2 x0 y0 z0 i x y z xi yi zi)))))))))). Time Defined. (* constant 1571 *) Definition l_e_st_eq_landau_n_rt_596_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_lessi (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0) (l_e_st_eq_landau_n_satz62c x y z (l_e_st_eq_landau_n_rt_lesse x0 y0 x y xix0 yiy0 l)))))))))))). Time Defined. (* constant 1572 *) Definition l_e_st_eq_landau_n_rt_satz96c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_596_t3 x0 y0 z0 l x y z xi yi zi)))))))))). Time Defined. (* constant 1573 *) Definition l_e_st_eq_landau_n_rt_596_andersa : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_satz95 x0 y0 z0 m)))). Time Defined. (* constant 1574 *) Definition l_e_st_eq_landau_n_rt_596_andersb : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ispl1 x0 y0 z0 i)))). Time Defined. (* constant 1575 *) Definition l_e_st_eq_landau_n_rt_596_andersc : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_satz95 y0 x0 z0 (l_e_st_eq_landau_n_rt_satz83 x0 y0 l)))))). Time Defined. (* constant 1576 *) Definition l_e_st_eq_landau_n_rt_satz96d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl z0 x0) (l_e_st_eq_landau_n_rt_pl z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_ismore12 (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl z0 x0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_pl z0 y0) (l_e_st_eq_landau_n_rt_compl x0 z0) (l_e_st_eq_landau_n_rt_compl y0 z0) (l_e_st_eq_landau_n_rt_satz96a x0 y0 z0 m))))). Time Defined. (* constant 1577 *) Definition l_e_st_eq_landau_n_rt_satz96e : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl z0 x0) (l_e_st_eq_landau_n_rt_pl z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ispl2 x0 y0 z0 i)))). Time Defined. (* constant 1578 *) Definition l_e_st_eq_landau_n_rt_satz96f : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl z0 x0) (l_e_st_eq_landau_n_rt_pl z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl z0 x0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_pl z0 y0) (l_e_st_eq_landau_n_rt_compl x0 z0) (l_e_st_eq_landau_n_rt_compl y0 z0) (l_e_st_eq_landau_n_rt_satz96c x0 y0 z0 l))))). Time Defined. (* constant 1579 *) Definition l_e_st_eq_landau_n_rt_597_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_more x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_morei x0 y0 x y xix0 yiy0 (l_e_st_eq_landau_n_satz63a x y z (l_e_st_eq_landau_n_rt_moree (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0) m)))))))))))). Time Defined. (* constant 1580 *) Definition l_e_st_eq_landau_n_rt_satz97a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_more x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_more x0 y0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_597_t1 x0 y0 z0 m x y z xi yi zi)))))))))). Time Defined. (* constant 1581 *) Definition l_e_st_eq_landau_n_rt_597_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_is x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_isi x0 y0 x y xix0 yiy0 (l_e_st_eq_landau_n_satz63b x y z (l_e_st_eq_landau_n_rt_ise (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0) i)))))))))))). Time Defined. (* constant 1582 *) Definition l_e_st_eq_landau_n_rt_satz97b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_is x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_is x0 y0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_597_t2 x0 y0 z0 i x y z xi yi zi)))))))))). Time Defined. (* constant 1583 *) Definition l_e_st_eq_landau_n_rt_597_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_less x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_lessi x0 y0 x y xix0 yiy0 (l_e_st_eq_landau_n_satz63c x y z (l_e_st_eq_landau_n_rt_lesse (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y z) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0) l)))))))))))). Time Defined. (* constant 1584 *) Definition l_e_st_eq_landau_n_rt_satz97c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_less x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_less x0 y0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_597_t3 x0 y0 z0 l x y z xi yi zi)))))))))). Time Defined. (* constant 1585 *) Definition l_e_st_eq_landau_n_rt_597_anders : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_less x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_satz82 y0 x0 (l_e_st_eq_landau_n_rt_satz97a y0 x0 z0 (l_e_st_eq_landau_n_rt_satz83 (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 z0) l)))))). Time Defined. (* constant 1586 *) Definition l_e_st_eq_landau_n_rt_598_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more z0 u0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more z0 u0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 u0 y u yiy0 uiu0) (l_e_st_eq_landau_n_satz64 x y z u (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m) (l_e_st_eq_landau_n_rt_moree z0 u0 z u ziz0 uiu0 n)))))))))))))))). Time Defined. (* constant 1587 *) Definition l_e_st_eq_landau_n_rt_satz98 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more z0 u0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more z0 u0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 z0 u0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_598_t1 x0 y0 z0 u0 m n x y z u xi yi zi ui)))))))))))))). Time Defined. (* constant 1588 *) Definition l_e_st_eq_landau_n_rt_satz98a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less z0 u0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less z0 u0) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_satz98 y0 x0 u0 z0 (l_e_st_eq_landau_n_rt_satz83 x0 y0 l) (l_e_st_eq_landau_n_rt_satz83 z0 u0 k)))))))). Time Defined. (* constant 1589 *) Definition l_e_st_eq_landau_n_rt_599_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more z0 u0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more z0 u0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 u0 y u yiy0 uiu0) (l_e_st_eq_landau_n_satz65a x y z u (l_e_st_eq_landau_n_rt_moreise x0 y0 x y xix0 yiy0 m) (l_e_st_eq_landau_n_rt_moree z0 u0 z u ziz0 uiu0 n)))))))))))))))). Time Defined. (* constant 1590 *) Definition l_e_st_eq_landau_n_rt_satz99a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more z0 u0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more z0 u0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 z0 u0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_599_t1 x0 y0 z0 u0 m n x y z u xi yi zi ui)))))))))))))). Time Defined. (* constant 1591 *) Definition l_e_st_eq_landau_n_rt_599_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis z0 u0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis z0 u0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 u0 y u yiy0 uiu0) (l_e_st_eq_landau_n_satz65b x y z u (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m) (l_e_st_eq_landau_n_rt_moreise z0 u0 z u ziz0 uiu0 n)))))))))))))))). Time Defined. (* constant 1592 *) Definition l_e_st_eq_landau_n_rt_satz99b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis z0 u0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis z0 u0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 z0 u0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_599_t2 x0 y0 z0 u0 m n x y z u xi yi zi ui)))))))))))))). Time Defined. (* constant 1593 *) Definition l_e_st_eq_landau_n_rt_satz99c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less z0 u0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less z0 u0) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_satz99a y0 x0 u0 z0 (l_e_st_eq_landau_n_rt_satz85 x0 y0 l) (l_e_st_eq_landau_n_rt_satz83 z0 u0 k)))))))). Time Defined. (* constant 1594 *) Definition l_e_st_eq_landau_n_rt_satz99d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis z0 u0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis z0 u0) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_satz99b y0 x0 u0 z0 (l_e_st_eq_landau_n_rt_satz83 x0 y0 l) (l_e_st_eq_landau_n_rt_satz85 z0 u0 k)))))))). Time Defined. (* constant 1595 *) Definition l_e_st_eq_landau_n_rt_5100_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis z0 u0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis z0 u0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_moreisi (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_pf x z) (l_e_st_eq_landau_n_pf y u) (l_e_st_eq_landau_n_rt_picp x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_picp y0 u0 y u yiy0 uiu0) (l_e_st_eq_landau_n_satz66 x y z u (l_e_st_eq_landau_n_rt_moreise x0 y0 x y xix0 yiy0 m) (l_e_st_eq_landau_n_rt_moreise z0 u0 z u ziz0 uiu0 n)))))))))))))))). Time Defined. (* constant 1596 *) Definition l_e_st_eq_landau_n_rt_satz100 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis z0 u0), l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis z0 u0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 z0 u0 (l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_5100_t1 x0 y0 z0 u0 m n x y z u xi yi zi ui)))))))))))))). Time Defined. (* constant 1597 *) Definition l_e_st_eq_landau_n_rt_satz100a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis z0 u0), l_e_st_eq_landau_n_rt_lessis (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_pl y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis z0 u0) => l_e_st_eq_landau_n_rt_satz84 (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_satz100 y0 x0 u0 z0 (l_e_st_eq_landau_n_rt_satz85 x0 y0 l) (l_e_st_eq_landau_n_rt_satz85 z0 u0 k)))))))). Time Defined. (* constant 1598 *) Definition l_e_st_eq_landau_n_rt_5101_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0), l_e_st_eq_landau_n_rt_more x0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0) => l_e_st_eq_landau_n_rt_ismore1 (l_e_st_eq_landau_n_rt_pl y0 v0) x0 y0 i (l_e_st_eq_landau_n_rt_satz94 y0 v0)))))). Time Defined. (* constant 1599 *) Definition l_e_st_eq_landau_n_rt_5101_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (v0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_nis (l_e_st_eq_landau_n_rt_pl y0 v0) x0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => l_imp_th3 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_satz81d x0 y0 l) (fun (t:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0) => l_e_st_eq_landau_n_rt_5101_t1 x0 y0 l v0 t))))). Time Defined. (* constant 1600 *) Definition l_e_st_eq_landau_n_rt_vorbemerkung101 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), l_not (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => l_some_th5 l_e_st_eq_landau_n_rt_rat (fun (v:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v) x0) (fun (v:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_5101_t2 x0 y0 l v)))). Time Defined. (* constant 1601 *) Definition l_e_st_eq_landau_n_rt_5101_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (v:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_ratof v)) x0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_ratof v)) x0 (l_e_st_eq_landau_n_pf y v) x (l_e_st_eq_landau_n_rt_picp y0 (l_e_st_eq_landau_n_rt_ratof v) y v yiy0 (l_e_st_eq_landau_n_rt_inclass v)) xix0 e))))))))). Time Defined. (* constant 1602 *) Definition l_e_st_eq_landau_n_rt_5101_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (v:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y v) x) => l_somei l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0) (l_e_st_eq_landau_n_rt_ratof v) (l_e_st_eq_landau_n_rt_5101_t3 x0 y0 m x y xix0 yiy0 v e)))))))))). Time Defined. (* constant 1603 *) Definition l_e_st_eq_landau_n_rt_5101_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_someapp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y t) x) (l_e_st_eq_landau_n_satz67a x y (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m)) (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0)) (fun (t:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf y t) x) => l_e_st_eq_landau_n_rt_5101_t4 x0 y0 m x y xix0 yiy0 t u))))))))). Time Defined. (* constant 1604 *) Definition l_e_st_eq_landau_n_rt_satz101a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_5101_t5 x0 y0 m x y xi yi))))))). Time Defined. (* constant 1605 *) Definition l_e_st_eq_landau_n_rt_5101_t6 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0), (forall (j:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 w0) x0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (v:l_e_st_eq_landau_n_frac), (forall (w:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (viv0:l_e_st_eq_landau_n_rt_inf v (l_e_st_eq_landau_n_rt_class v0)), (forall (wiw0:l_e_st_eq_landau_n_rt_inf w (l_e_st_eq_landau_n_rt_class w0)), l_e_st_eq_landau_n_rt_is v0 w0)))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0) => (fun (j:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 w0) x0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (w:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (viv0:l_e_st_eq_landau_n_rt_inf v (l_e_st_eq_landau_n_rt_class v0)) => (fun (wiw0:l_e_st_eq_landau_n_rt_inf w (l_e_st_eq_landau_n_rt_class w0)) => l_e_st_eq_landau_n_rt_isi v0 w0 v w viv0 wiw0 (l_e_st_eq_landau_n_satz67b x y v w (l_e_st_eq_landau_n_rt_ise (l_e_st_eq_landau_n_rt_pl y0 v0) x0 (l_e_st_eq_landau_n_pf y v) x (l_e_st_eq_landau_n_rt_picp y0 v0 y v yiy0 viv0) xix0 i) (l_e_st_eq_landau_n_rt_ise (l_e_st_eq_landau_n_rt_pl y0 w0) x0 (l_e_st_eq_landau_n_pf y w) x (l_e_st_eq_landau_n_rt_picp y0 w0 y w yiy0 wiw0) xix0 j)))))))))))))))). Time Defined. (* constant 1606 *) Definition l_e_st_eq_landau_n_rt_satz101b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0), (forall (j:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 w0) x0), l_e_st_eq_landau_n_rt_is v0 w0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0) => (fun (j:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 w0) x0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 v0 w0 (l_e_st_eq_landau_n_rt_is v0 w0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (w:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (vi:l_e_st_eq_landau_n_rt_inf v (l_e_st_eq_landau_n_rt_class v0)) => (fun (wi:l_e_st_eq_landau_n_rt_inf w (l_e_st_eq_landau_n_rt_class w0)) => l_e_st_eq_landau_n_rt_5101_t6 x0 y0 v0 w0 i j x y v w xi yi vi wi)))))))))))))). Time Defined. (* constant 1607 *) Definition l_e_st_eq_landau_n_rt_5101_t7 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_amone l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_rat) => (fun (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0) => (fun (w:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 u) x0) => l_e_st_eq_landau_n_rt_satz101b x0 y0 t u v w)))))). Time Defined. (* constant 1608 *) Definition l_e_st_eq_landau_n_rt_satz101 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_one (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_onei l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0) (l_e_st_eq_landau_n_rt_5101_t7 x0 y0) (l_e_st_eq_landau_n_rt_satz101a x0 y0 m)))). Time Defined. (* constant 1609 *) Definition l_e_st_eq_landau_n_rt_mn : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_rat))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_ind l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0) (l_e_st_eq_landau_n_rt_satz101 x0 y0 m)))). Time Defined. (* constant 1610 *) Definition l_e_st_eq_landau_n_rt_satz101c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_oneax l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 t) x0) (l_e_st_eq_landau_n_rt_satz101 x0 y0 m)))). Time Defined. (* constant 1611 *) Definition l_e_st_eq_landau_n_rt_satz101d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_mn x0 y0 m))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) x0 (l_e_st_eq_landau_n_rt_satz101c x0 y0 m)))). Time Defined. (* constant 1612 *) Definition l_e_st_eq_landau_n_rt_satz101e : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0) (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) x0 (l_e_st_eq_landau_n_rt_compl (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0) (l_e_st_eq_landau_n_rt_satz101c x0 y0 m)))). Time Defined. (* constant 1613 *) Definition l_e_st_eq_landau_n_rt_satz101f : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0) x0 (l_e_st_eq_landau_n_rt_satz101e x0 y0 m)))). Time Defined. (* constant 1614 *) Definition l_e_st_eq_landau_n_rt_satz101g : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0), l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 v0) x0) => l_e_st_eq_landau_n_rt_satz101b x0 y0 v0 (l_e_st_eq_landau_n_rt_mn x0 y0 m) i (l_e_st_eq_landau_n_rt_satz101c x0 y0 m)))))). Time Defined. (* constant 1615 *) Definition l_e_st_eq_landau_n_rt_timesfrt : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_rt_rat)). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_tf x y))). Time Defined. (* constant 1616 *) Definition l_e_st_eq_landau_n_rt_ii5_t20 : (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq x y), (forall (f:l_e_st_eq_landau_n_eq z u), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_timesfrt x z) (l_e_st_eq_landau_n_rt_timesfrt y u))))))). exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq x y) => (fun (f:l_e_st_eq_landau_n_eq z u) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_tf x z)) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_tf y u)) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_tf x z)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_tf y u)) (l_e_st_eq_landau_n_satz68 x y z u e f))))))). Time Defined. (* constant 1617 *) Definition l_e_st_eq_landau_n_rt_ftimesfrt : l_e_st_eq_landau_n_rt_fixf l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_timesfrt. exact (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_rt_eq x y) => (fun (w:l_e_st_eq_landau_n_rt_eq z u) => l_e_st_eq_landau_n_rt_ii5_t20 x y z u v w)))))). Time Defined. (* constant 1618 *) Definition l_e_st_eq_landau_n_rt_ts : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rat)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_indrat x0 y0 l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_timesfrt l_e_st_eq_landau_n_rt_ftimesfrt)). Time Defined. (* constant 1619 *) Definition l_e_st_eq_landau_n_rt_ii5_t21 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_tf x1 y1)) (l_e_st_eq_landau_n_rt_ts x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_isindrat x0 y0 l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_timesfrt l_e_st_eq_landau_n_rt_ftimesfrt x1 y1 x1ix0 y1iy0)))))). Time Defined. (* constant 1620 *) Definition l_e_st_eq_landau_n_rt_tict : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_tf x1 y1) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_ts x0 y0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_tf x1 y1) (l_e_st_eq_landau_n_rt_class t)) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_tf x1 y1)) (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_tf x1 y1)) (l_e_st_eq_landau_n_rt_ii5_t21 x0 y0 x1 y1 x1ix0 y1iy0))))))). Time Defined. (* constant 1621 *) Definition l_e_st_eq_landau_n_rt_ists1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_isf l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ts t z0) x0 y0 i)))). Time Defined. (* constant 1622 *) Definition l_e_st_eq_landau_n_rt_ists2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts z0 x0) (l_e_st_eq_landau_n_rt_ts z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_isf l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ts z0 t) x0 y0 i)))). Time Defined. (* constant 1623 *) Definition l_e_st_eq_landau_n_rt_ists12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 u0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 u0) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_rt_ists1 x0 y0 z0 i) (l_e_st_eq_landau_n_rt_ists2 z0 u0 y0 j))))))). Time Defined. (* constant 1624 *) Definition l_e_st_eq_landau_n_rt_5102_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts y0 x0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts y0 x0) (l_e_st_eq_landau_n_tf x1 y1) (l_e_st_eq_landau_n_tf y1 x1) (l_e_st_eq_landau_n_rt_tict x0 y0 x1 y1 x1ix0 y1iy0) (l_e_st_eq_landau_n_rt_tict y0 x0 y1 x1 y1iy0 x1ix0) (l_e_st_eq_landau_n_satz69 x1 y1))))))). Time Defined. (* constant 1625 *) Definition l_e_st_eq_landau_n_rt_satz102 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts y0 x0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_5102_t1 x0 y0 x y xi yi)))))). Time Defined. (* constant 1626 *) Definition l_e_st_eq_landau_n_rt_comts : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts y0 x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz102 x0 y0)). Time Defined. (* constant 1627 *) Definition l_e_st_eq_landau_n_rt_5103_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_tf x y) z) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_tict (l_e_st_eq_landau_n_rt_ts x0 y0) z0 (l_e_st_eq_landau_n_tf x y) z (l_e_st_eq_landau_n_rt_tict x0 y0 x y xix0 yiy0) ziz0))))))))). Time Defined. (* constant 1628 *) Definition l_e_st_eq_landau_n_rt_5103_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0)))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_tict x0 (l_e_st_eq_landau_n_rt_ts y0 z0) x (l_e_st_eq_landau_n_tf y z) xix0 (l_e_st_eq_landau_n_rt_tict y0 z0 y z yiy0 ziz0)))))))))). Time Defined. (* constant 1629 *) Definition l_e_st_eq_landau_n_rt_5103_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0)) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_tf x y) z) (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_tf y z)) (l_e_st_eq_landau_n_rt_5103_t1 x0 y0 z0 x y z xix0 yiy0 ziz0) (l_e_st_eq_landau_n_rt_5103_t2 x0 y0 z0 x y z xix0 yiy0 ziz0) (l_e_st_eq_landau_n_satz70 x y z)))))))))). Time Defined. (* constant 1630 *) Definition l_e_st_eq_landau_n_rt_satz103 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0))) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_5103_t3 x0 y0 z0 x y z xi yi zi))))))))). Time Defined. (* constant 1631 *) Definition l_e_st_eq_landau_n_rt_assts1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz103 x0 y0 z0))). Time Defined. (* constant 1632 *) Definition l_e_st_eq_landau_n_rt_assts2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0)) (l_e_st_eq_landau_n_rt_satz103 x0 y0 z0)))). Time Defined. (* constant 1633 *) Definition l_e_st_eq_landau_n_rt_5104_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0)))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_tict x0 (l_e_st_eq_landau_n_rt_pl y0 z0) x (l_e_st_eq_landau_n_pf y z) xix0 (l_e_st_eq_landau_n_rt_picp y0 z0 y z yiy0 ziz0)))))))))). Time Defined. (* constant 1634 *) Definition l_e_st_eq_landau_n_rt_5104_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf x z)) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0)))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_picp (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_rt_tict x0 y0 x y xix0 yiy0) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0)))))))))). Time Defined. (* constant 1635 *) Definition l_e_st_eq_landau_n_rt_5104_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0)) (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_pf y z)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_tf x y) (l_e_st_eq_landau_n_tf x z)) (l_e_st_eq_landau_n_rt_5104_t1 x0 y0 z0 x y z xix0 yiy0 ziz0) (l_e_st_eq_landau_n_rt_5104_t2 x0 y0 z0 x y z xix0 yiy0 ziz0) (l_e_st_eq_landau_n_satz71 x y z)))))))))). Time Defined. (* constant 1636 *) Definition l_e_st_eq_landau_n_rt_satz104 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0))) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_5104_t3 x0 y0 z0 x y z xi yi zi))))))))). Time Defined. (* constant 1637 *) Definition l_e_st_eq_landau_n_rt_disttp1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_ts z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts z0 x0) (l_e_st_eq_landau_n_rt_ts z0 y0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) (l_e_st_eq_landau_n_rt_comts (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_satz104 z0 x0 y0) (l_e_st_eq_landau_n_rt_ispl12 (l_e_st_eq_landau_n_rt_ts z0 x0) (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts z0 y0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_rt_comts z0 x0) (l_e_st_eq_landau_n_rt_comts z0 y0))))). Time Defined. (* constant 1638 *) Definition l_e_st_eq_landau_n_rt_disttp2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz104 x0 y0 z0))). Time Defined. (* constant 1639 *) Definition l_e_st_eq_landau_n_rt_distpt1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_pl x0 y0) z0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) (l_e_st_eq_landau_n_rt_disttp1 x0 y0 z0)))). Time Defined. (* constant 1640 *) Definition l_e_st_eq_landau_n_rt_distpt2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0)) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0)) (l_e_st_eq_landau_n_rt_disttp2 x0 y0 z0)))). Time Defined. (* constant 1641 *) Definition l_e_st_eq_landau_n_rt_5105_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 z0 y z yiy0 ziz0) (l_e_st_eq_landau_n_satz72a x y z (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m)))))))))))). Time Defined. (* constant 1642 *) Definition l_e_st_eq_landau_n_rt_satz105a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_5105_t1 x0 y0 z0 m x y z xi yi zi)))))))))). Time Defined. (* constant 1643 *) Definition l_e_st_eq_landau_n_rt_5105_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 z0 y z yiy0 ziz0) (l_e_st_eq_landau_n_satz72b x y z (l_e_st_eq_landau_n_rt_ise x0 y0 x y xix0 yiy0 i)))))))))))). Time Defined. (* constant 1644 *) Definition l_e_st_eq_landau_n_rt_satz105b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_5105_t2 x0 y0 z0 i x y z xi yi zi)))))))))). Time Defined. (* constant 1645 *) Definition l_e_st_eq_landau_n_rt_5105_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_lessi (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xi zi) (l_e_st_eq_landau_n_rt_tict y0 z0 y z yi zi) (l_e_st_eq_landau_n_satz72c x y z (l_e_st_eq_landau_n_rt_lesse x0 y0 x y xi yi l)))))))))))). Time Defined. (* constant 1646 *) Definition l_e_st_eq_landau_n_rt_satz105c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_5105_t3 x0 y0 z0 l x y z xi yi zi)))))))))). Time Defined. (* constant 1647 *) Definition l_e_st_eq_landau_n_rt_5105_andersb : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ists1 x0 y0 z0 i)))). Time Defined. (* constant 1648 *) Definition l_e_st_eq_landau_n_rt_5105_andersc : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_satz105a y0 x0 z0 (l_e_st_eq_landau_n_rt_satz83 x0 y0 l)))))). Time Defined. (* constant 1649 *) Definition l_e_st_eq_landau_n_rt_satz105d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts z0 x0) (l_e_st_eq_landau_n_rt_ts z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_ismore12 (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts z0 x0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_rt_ts z0 y0) (l_e_st_eq_landau_n_rt_comts x0 z0) (l_e_st_eq_landau_n_rt_comts y0 z0) (l_e_st_eq_landau_n_rt_satz105a x0 y0 z0 m))))). Time Defined. (* constant 1650 *) Definition l_e_st_eq_landau_n_rt_satz105e : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts z0 x0) (l_e_st_eq_landau_n_rt_ts z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ists2 x0 y0 z0 i)))). Time Defined. (* constant 1651 *) Definition l_e_st_eq_landau_n_rt_satz105f : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts z0 x0) (l_e_st_eq_landau_n_rt_ts z0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts z0 x0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_rt_ts z0 y0) (l_e_st_eq_landau_n_rt_comts x0 z0) (l_e_st_eq_landau_n_rt_comts y0 z0) (l_e_st_eq_landau_n_rt_satz105c x0 y0 z0 l))))). Time Defined. (* constant 1652 *) Definition l_e_st_eq_landau_n_rt_5106_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_more x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_morei x0 y0 x y xix0 yiy0 (l_e_st_eq_landau_n_satz73a x y z (l_e_st_eq_landau_n_rt_moree (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 z0 y z yiy0 ziz0) m)))))))))))). Time Defined. (* constant 1653 *) Definition l_e_st_eq_landau_n_rt_satz106a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)), l_e_st_eq_landau_n_rt_more x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_more x0 y0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_5106_t1 x0 y0 z0 m x y z xi yi zi)))))))))). Time Defined. (* constant 1654 *) Definition l_e_st_eq_landau_n_rt_5106_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_is x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_isi x0 y0 x y xix0 yiy0 (l_e_st_eq_landau_n_satz73b x y z (l_e_st_eq_landau_n_rt_ise (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 z0 y z yiy0 ziz0) i)))))))))))). Time Defined. (* constant 1655 *) Definition l_e_st_eq_landau_n_rt_satz106b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)), l_e_st_eq_landau_n_rt_is x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_is x0 y0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_5106_t2 x0 y0 z0 i x y z xi yi zi)))))))))). Time Defined. (* constant 1656 *) Definition l_e_st_eq_landau_n_rt_5106_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), l_e_st_eq_landau_n_rt_less x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_lessi x0 y0 x y xix0 yiy0 (l_e_st_eq_landau_n_satz73c x y z (l_e_st_eq_landau_n_rt_lesse (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y z) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 z0 y z yiy0 ziz0) l)))))))))))). Time Defined. (* constant 1657 *) Definition l_e_st_eq_landau_n_rt_satz106c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)), l_e_st_eq_landau_n_rt_less x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) => l_e_st_eq_landau_n_rt_ratapp3 x0 y0 z0 (l_e_st_eq_landau_n_rt_less x0 y0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => l_e_st_eq_landau_n_rt_5106_t3 x0 y0 z0 l x y z xi yi zi)))))))))). Time Defined. (* constant 1658 *) Definition l_e_st_eq_landau_n_rt_5106_anders : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)), l_e_st_eq_landau_n_rt_less x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0)) => l_e_st_eq_landau_n_rt_satz82 y0 x0 (l_e_st_eq_landau_n_rt_satz106a y0 x0 z0 (l_e_st_eq_landau_n_rt_satz83 (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 z0) l)))))). Time Defined. (* constant 1659 *) Definition l_e_st_eq_landau_n_rt_5107_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more z0 u0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more z0 u0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 u0 y u yiy0 uiu0) (l_e_st_eq_landau_n_satz74 x y z u (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m) (l_e_st_eq_landau_n_rt_moree z0 u0 z u ziz0 uiu0 n)))))))))))))))). Time Defined. (* constant 1660 *) Definition l_e_st_eq_landau_n_rt_satz107 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more z0 u0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more z0 u0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 z0 u0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_5107_t1 x0 y0 z0 u0 m n x y z u xi yi zi ui)))))))))))))). Time Defined. (* constant 1661 *) Definition l_e_st_eq_landau_n_rt_satz107a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less z0 u0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less z0 u0) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_satz107 y0 x0 u0 z0 (l_e_st_eq_landau_n_rt_satz83 x0 y0 l) (l_e_st_eq_landau_n_rt_satz83 z0 u0 k)))))))). Time Defined. (* constant 1662 *) Definition l_e_st_eq_landau_n_rt_5108_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more z0 u0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more z0 u0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 u0 y u yiy0 uiu0) (l_e_st_eq_landau_n_satz75a x y z u (l_e_st_eq_landau_n_rt_moreise x0 y0 x y xix0 yiy0 m) (l_e_st_eq_landau_n_rt_moree z0 u0 z u ziz0 uiu0 n)))))))))))))))). Time Defined. (* constant 1663 *) Definition l_e_st_eq_landau_n_rt_satz108a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_more z0 u0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_more z0 u0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 z0 u0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_5108_t1 x0 y0 z0 u0 m n x y z u xi yi zi ui)))))))))))))). Time Defined. (* constant 1664 *) Definition l_e_st_eq_landau_n_rt_5108_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis z0 u0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis z0 u0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 u0 y u yiy0 uiu0) (l_e_st_eq_landau_n_satz75b x y z u (l_e_st_eq_landau_n_rt_moree x0 y0 x y xix0 yiy0 m) (l_e_st_eq_landau_n_rt_moreise z0 u0 z u ziz0 uiu0 n)))))))))))))))). Time Defined. (* constant 1665 *) Definition l_e_st_eq_landau_n_rt_satz108b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis z0 u0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis z0 u0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 z0 u0 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_5108_t2 x0 y0 z0 u0 m n x y z u xi yi zi ui)))))))))))))). Time Defined. (* constant 1666 *) Definition l_e_st_eq_landau_n_rt_satz108c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_less z0 u0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_less z0 u0) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_satz108a y0 x0 u0 z0 (l_e_st_eq_landau_n_rt_satz85 x0 y0 l) (l_e_st_eq_landau_n_rt_satz83 z0 u0 k)))))))). Time Defined. (* constant 1667 *) Definition l_e_st_eq_landau_n_rt_satz108d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis z0 u0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis z0 u0) => l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_satz108b y0 x0 u0 z0 (l_e_st_eq_landau_n_rt_satz83 x0 y0 l) (l_e_st_eq_landau_n_rt_satz85 z0 u0 k)))))))). Time Defined. (* constant 1668 *) Definition l_e_st_eq_landau_n_rt_5109_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis z0 u0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (z:l_e_st_eq_landau_n_frac), (forall (u:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis z0 u0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (ziz0:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_moreisi (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_tf x z) (l_e_st_eq_landau_n_tf y u) (l_e_st_eq_landau_n_rt_tict x0 z0 x z xix0 ziz0) (l_e_st_eq_landau_n_rt_tict y0 u0 y u yiy0 uiu0) (l_e_st_eq_landau_n_satz76 x y z u (l_e_st_eq_landau_n_rt_moreise x0 y0 x y xix0 yiy0 m) (l_e_st_eq_landau_n_rt_moreise z0 u0 z u ziz0 uiu0 n)))))))))))))))). Time Defined. (* constant 1669 *) Definition l_e_st_eq_landau_n_rt_satz109 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), (forall (n:l_e_st_eq_landau_n_rt_moreis z0 u0), l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => (fun (n:l_e_st_eq_landau_n_rt_moreis z0 u0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 z0 u0 (l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (z:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (zi:l_e_st_eq_landau_n_rt_inf z (l_e_st_eq_landau_n_rt_class z0)) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_5109_t1 x0 y0 z0 u0 m n x y z u xi yi zi ui)))))))))))))). Time Defined. (* constant 1670 *) Definition l_e_st_eq_landau_n_rt_satz109a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), (forall (k:l_e_st_eq_landau_n_rt_lessis z0 u0), l_e_st_eq_landau_n_rt_lessis (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_ts y0 u0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => (fun (k:l_e_st_eq_landau_n_rt_lessis z0 u0) => l_e_st_eq_landau_n_rt_satz84 (l_e_st_eq_landau_n_rt_ts y0 u0) (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_satz109 y0 x0 u0 z0 (l_e_st_eq_landau_n_rt_satz85 x0 y0 l) (l_e_st_eq_landau_n_rt_satz85 z0 u0 k)))))))). Time Defined. (* constant 1671 *) Definition l_e_st_eq_landau_n_rt_5110_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (v:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y1 v) x1), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ratof v)) x0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y1 v) x1) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ratof v)) x0 (l_e_st_eq_landau_n_tf y1 v) x1 (l_e_st_eq_landau_n_rt_tict y0 (l_e_st_eq_landau_n_rt_ratof v) y1 v y1iy0 (l_e_st_eq_landau_n_rt_inclass v)) x1ix0 e)))))))). Time Defined. (* constant 1672 *) Definition l_e_st_eq_landau_n_rt_5110_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), (forall (v:l_e_st_eq_landau_n_frac), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y1 v) x1), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y1 v) x1) => l_somei l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0) (l_e_st_eq_landau_n_rt_ratof v) (l_e_st_eq_landau_n_rt_5110_t1 x0 y0 x1 y1 x1ix0 y1iy0 v e))))))))). Time Defined. (* constant 1673 *) Definition l_e_st_eq_landau_n_rt_5110_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x1:l_e_st_eq_landau_n_frac), (forall (y1:l_e_st_eq_landau_n_frac), (forall (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)), (forall (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x1:l_e_st_eq_landau_n_frac) => (fun (y1:l_e_st_eq_landau_n_frac) => (fun (x1ix0:l_e_st_eq_landau_n_rt_inf x1 (l_e_st_eq_landau_n_rt_class x0)) => (fun (y1iy0:l_e_st_eq_landau_n_rt_inf y1 (l_e_st_eq_landau_n_rt_class y0)) => l_someapp l_e_st_eq_landau_n_frac (fun (t:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y1 t) x1) (l_e_st_eq_landau_n_satz77a x1 y1) (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0)) (fun (t:l_e_st_eq_landau_n_frac) => (fun (u:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf y1 t) x1) => l_e_st_eq_landau_n_rt_5110_t2 x0 y0 x1 y1 x1ix0 y1iy0 t u)))))))). Time Defined. (* constant 1674 *) Definition l_e_st_eq_landau_n_rt_satz110a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp2 x0 y0 (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0)) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => l_e_st_eq_landau_n_rt_5110_t3 x0 y0 x y xi yi)))))). Time Defined. (* constant 1675 *) Definition l_e_st_eq_landau_n_rt_5110_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 v0) x0), (forall (j:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 w0) x0), (forall (x:l_e_st_eq_landau_n_frac), (forall (y:l_e_st_eq_landau_n_frac), (forall (v:l_e_st_eq_landau_n_frac), (forall (w:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), (forall (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)), (forall (viv0:l_e_st_eq_landau_n_rt_inf v (l_e_st_eq_landau_n_rt_class v0)), (forall (wiw0:l_e_st_eq_landau_n_rt_inf w (l_e_st_eq_landau_n_rt_class w0)), l_e_st_eq_landau_n_rt_is v0 w0)))))))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 v0) x0) => (fun (j:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 w0) x0) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (w:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yiy0:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (viv0:l_e_st_eq_landau_n_rt_inf v (l_e_st_eq_landau_n_rt_class v0)) => (fun (wiw0:l_e_st_eq_landau_n_rt_inf w (l_e_st_eq_landau_n_rt_class w0)) => l_e_st_eq_landau_n_rt_isi v0 w0 v w viv0 wiw0 (l_e_st_eq_landau_n_satz77b x y v w (l_e_st_eq_landau_n_rt_ise (l_e_st_eq_landau_n_rt_ts y0 v0) x0 (l_e_st_eq_landau_n_tf y v) x (l_e_st_eq_landau_n_rt_tict y0 v0 y v yiy0 viv0) xix0 i) (l_e_st_eq_landau_n_rt_ise (l_e_st_eq_landau_n_rt_ts y0 w0) x0 (l_e_st_eq_landau_n_tf y w) x (l_e_st_eq_landau_n_rt_tict y0 w0 y w yiy0 wiw0) xix0 j)))))))))))))))). Time Defined. (* constant 1676 *) Definition l_e_st_eq_landau_n_rt_satz110b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 v0) x0), (forall (j:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 w0) x0), l_e_st_eq_landau_n_rt_is v0 w0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 v0) x0) => (fun (j:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 w0) x0) => l_e_st_eq_landau_n_rt_ratapp4 x0 y0 v0 w0 (l_e_st_eq_landau_n_rt_is v0 w0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (y:l_e_st_eq_landau_n_frac) => (fun (v:l_e_st_eq_landau_n_frac) => (fun (w:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => (fun (yi:l_e_st_eq_landau_n_rt_inf y (l_e_st_eq_landau_n_rt_class y0)) => (fun (vi:l_e_st_eq_landau_n_rt_inf v (l_e_st_eq_landau_n_rt_class v0)) => (fun (wi:l_e_st_eq_landau_n_rt_inf w (l_e_st_eq_landau_n_rt_class w0)) => l_e_st_eq_landau_n_rt_5110_t4 x0 y0 v0 w0 i j x y v w xi yi vi wi)))))))))))))). Time Defined. (* constant 1677 *) Definition l_e_st_eq_landau_n_rt_5110_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_amone l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_rat) => (fun (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0) => (fun (w:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 u) x0) => l_e_st_eq_landau_n_rt_satz110b x0 y0 t u v w)))))). Time Defined. (* constant 1678 *) Definition l_e_st_eq_landau_n_rt_satz110 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_one (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_onei l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0) (l_e_st_eq_landau_n_rt_5110_t5 x0 y0) (l_e_st_eq_landau_n_rt_satz110a x0 y0))). Time Defined. (* constant 1679 *) Definition l_e_st_eq_landau_n_5111_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) x)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_ts x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_ndis12 x l_e_st_eq_landau_n_1 y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz28a x))). Time Defined. (* constant 1680 *) Definition l_e_st_eq_landau_n_5111_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) x (l_e_st_eq_landau_n_5111_t1 x y))). Time Defined. (* constant 1681 *) Definition l_e_st_eq_landau_n_satz111a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_more x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1))) y (l_e_st_eq_landau_n_5111_t1 x y) (l_e_st_eq_landau_n_5111_t1 y x) m))). Time Defined. (* constant 1682 *) Definition l_e_st_eq_landau_n_satz111b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_is x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (e:l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) => l_e_tr3is l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1))) y (l_e_st_eq_landau_n_5111_t2 x y) e (l_e_st_eq_landau_n_5111_t1 y x)))). Time Defined. (* constant 1683 *) Definition l_e_st_eq_landau_n_satz111c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_less x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1))) y (l_e_st_eq_landau_n_5111_t1 x y) (l_e_st_eq_landau_n_5111_t1 y x) l))). Time Defined. (* constant 1684 *) Definition l_e_st_eq_landau_n_satz111d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_ismore12 x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) y (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_5111_t2 x y) (l_e_st_eq_landau_n_5111_t2 y x) m))). Time Defined. (* constant 1685 *) Definition l_e_st_eq_landau_n_satz111e : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) x y (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_5111_t1 x y) i (l_e_st_eq_landau_n_5111_t2 y x)))). Time Defined. (* constant 1686 *) Definition l_e_st_eq_landau_n_satz111f : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_isless12 x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) y (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_5111_t2 x y) (l_e_st_eq_landau_n_5111_t2 y x) l))). Time Defined. (* constant 1687 *) Definition l_e_st_eq_landau_n_rt_natprop : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_class x0))). Time Defined. (* constant 1688 *) Definition l_e_st_eq_landau_n_rt_natrt : (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_some (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop x0 t)). Time Defined. (* constant 1689 *) Definition l_e_st_eq_landau_n_rt_ii5_t22 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (npx:l_e_st_eq_landau_n_rt_natprop x0 x), (forall (npy:l_e_st_eq_landau_n_rt_natprop y0 y), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_is x y))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (npx:l_e_st_eq_landau_n_rt_natprop x0 x) => (fun (npy:l_e_st_eq_landau_n_rt_natprop y0 y) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_satz111b x y (l_e_st_eq_landau_n_rt_ise x0 y0 (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1) npx npy i)))))))). Time Defined. (* constant 1690 *) Definition l_e_st_eq_landau_n_rt_ii5_t23 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_amone l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop x0 t)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_rt_natprop x0 t) => (fun (w:l_e_st_eq_landau_n_rt_natprop x0 u) => l_e_st_eq_landau_n_rt_ii5_t22 x0 x0 t u v w (l_e_refis l_e_st_eq_landau_n_rt_rat x0)))))). Time Defined. (* constant 1691 *) Definition l_e_st_eq_landau_n_rt_satz111g : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), l_e_st_eq_landau_n_one (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop x0 t))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => l_e_onei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop x0 t) (l_e_st_eq_landau_n_rt_ii5_t23 x0) nx0)). Time Defined. (* constant 1692 *) Definition l_e_st_eq_landau_n_rt_nofrt : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), l_e_st_eq_landau_n_nat)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => l_e_ind l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop x0 t) (l_e_st_eq_landau_n_rt_satz111g x0 nx0))). Time Defined. (* constant 1693 *) Definition l_e_st_eq_landau_n_rt_inclassn : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_class x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => l_e_oneax l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop x0 t) (l_e_st_eq_landau_n_rt_satz111g x0 nx0))). Time Defined. (* constant 1694 *) Definition l_e_st_eq_landau_n_rt_isrten : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_ii5_t22 x0 y0 (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) (l_e_st_eq_landau_n_rt_inclassn x0 nx0) (l_e_st_eq_landau_n_rt_inclassn y0 ny0) i))))). Time Defined. (* constant 1695 *) Definition l_e_st_eq_landau_n_rt_isrtin : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)), l_e_st_eq_landau_n_rt_is x0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)) => l_e_st_eq_landau_n_rt_isi x0 y0 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt y0 ny0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn x0 nx0) (l_e_st_eq_landau_n_rt_inclassn y0 ny0) (l_e_st_eq_landau_n_eqn (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_nofrt y0 ny0) i)))))). Time Defined. (* constant 1696 *) Definition l_e_st_eq_landau_n_rt_rtofn : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rat). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)). Time Defined. (* constant 1697 *) Definition l_e_st_eq_landau_n_rt_natrti : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_rtofn x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop (l_e_st_eq_landau_n_rt_rtofn x) t) x (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1))). Time Defined. (* constant 1698 *) Definition l_e_st_eq_landau_n_rt_isnert : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rtofn t) x y i))). Time Defined. (* constant 1699 *) Definition l_e_st_eq_landau_n_rt_isnirt : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)), l_e_st_eq_landau_n_is x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) => l_e_st_eq_landau_n_rt_ii5_t22 (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y) x y (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) i))). Time Defined. (* constant 1700 *) Definition l_e_st_eq_landau_n_rt_isrtn1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_nofrt x0 nx0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => l_e_st_eq_landau_n_rt_isi x0 (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_nofrt x0 nx0)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn x0 nx0) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_refeq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1)))). Time Defined. (* constant 1701 *) Definition l_e_st_eq_landau_n_rt_isnrt1 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_ii5_t22 (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn x) x (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x)) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rtofn x))). Time Defined. (* constant 1702 *) Definition l_e_st_eq_landau_n_satz112a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl x y) l_e_st_eq_landau_n_1))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz57 x y l_e_st_eq_landau_n_1)). Time Defined. (* constant 1703 *) Definition l_e_st_eq_landau_n_satz112b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x y) l_e_st_eq_landau_n_1))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_treq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x y) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_tfeq12a x l_e_st_eq_landau_n_1 y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_eqd (l_e_st_eq_landau_n_ts x y) (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz28a l_e_st_eq_landau_n_1)))). Time Defined. (* constant 1704 *) Definition l_e_st_eq_landau_n_rt_satz112c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_pl x0 y0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => l_e_st_eq_landau_n_rt_lemmaeq1 (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt y0 ny0) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_picp x0 y0 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt y0 ny0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn x0 nx0) (l_e_st_eq_landau_n_rt_inclassn y0 ny0)) (l_e_st_eq_landau_n_satz112a (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)))))). Time Defined. (* constant 1705 *) Definition l_e_st_eq_landau_n_rt_satz112d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_pl x0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop (l_e_st_eq_landau_n_rt_pl x0 y0) t) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)) (l_e_st_eq_landau_n_rt_satz112c x0 nx0 y0 ny0))))). Time Defined. (* constant 1706 *) Definition l_e_st_eq_landau_n_rt_satz112e : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_ts x0 y0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => l_e_st_eq_landau_n_rt_lemmaeq1 (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt y0 ny0) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_tict x0 y0 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt y0 ny0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn x0 nx0) (l_e_st_eq_landau_n_rt_inclassn y0 ny0)) (l_e_st_eq_landau_n_satz112b (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)))))). Time Defined. (* constant 1707 *) Definition l_e_st_eq_landau_n_rt_satz112f : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_ts x0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_natprop (l_e_st_eq_landau_n_rt_ts x0 y0) t) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)) (l_e_st_eq_landau_n_rt_satz112e x0 nx0 y0 ny0))))). Time Defined. (* constant 1708 *) Definition l_e_st_eq_landau_n_rt_5112_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_satz111a (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) (l_e_st_eq_landau_n_rt_moree x0 y0 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt y0 ny0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn x0 nx0) (l_e_st_eq_landau_n_rt_inclassn y0 ny0) m)))))). Time Defined. (* constant 1709 *) Definition l_e_st_eq_landau_n_rt_5112_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (z:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nofrt y0 ny0) (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn z) (l_e_st_eq_landau_n_rt_natrti z)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nofrt y0 ny0) (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn z) (l_e_st_eq_landau_n_rt_natrti z))) d (l_e_st_eq_landau_n_ispl2 z (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn z) (l_e_st_eq_landau_n_rt_natrti z)) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) (l_e_st_eq_landau_n_rt_isnrt1 z))))))))). Time Defined. (* constant 1710 *) Definition l_e_st_eq_landau_n_rt_5112_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (z:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_rtofn z))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z) => l_e_st_eq_landau_n_rt_isi x0 (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_rtofn z)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nofrt y0 ny0) (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn z) (l_e_st_eq_landau_n_rt_natrti z))) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn x0 nx0) (l_e_st_eq_landau_n_rt_satz112c y0 ny0 (l_e_st_eq_landau_n_rt_rtofn z) (l_e_st_eq_landau_n_rt_natrti z)) (l_e_st_eq_landau_n_eqn (l_e_st_eq_landau_n_rt_nofrt x0 nx0) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nofrt y0 ny0) (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn z) (l_e_st_eq_landau_n_rt_natrti z))) (l_e_st_eq_landau_n_rt_5112_t2 x0 nx0 y0 ny0 m z d))))))))). Time Defined. (* constant 1711 *) Definition l_e_st_eq_landau_n_rt_5112_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (z:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn z) (l_e_st_eq_landau_n_rt_mn x0 y0 m)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z) => l_e_st_eq_landau_n_rt_satz101g x0 y0 (l_e_st_eq_landau_n_rt_rtofn z) m (l_e_symis l_e_st_eq_landau_n_rt_rat x0 (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_rtofn z)) (l_e_st_eq_landau_n_rt_5112_t3 x0 nx0 y0 ny0 m z d))))))))). Time Defined. (* constant 1712 *) Definition l_e_st_eq_landau_n_rt_5112_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (z:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_mn x0 y0 m)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (z:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) z) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) (l_e_st_eq_landau_n_rt_rtofn z) (l_e_st_eq_landau_n_rt_mn x0 y0 m) (l_e_st_eq_landau_n_rt_natrti z) (l_e_st_eq_landau_n_rt_5112_t4 x0 nx0 y0 ny0 m z d)))))))). Time Defined. (* constant 1713 *) Definition l_e_st_eq_landau_n_rt_satz112g : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_mn x0 y0 m)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_someapp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) t) (l_e_st_eq_landau_n_rt_5112_t1 x0 nx0 y0 ny0 m) (l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_mn x0 y0 m)) (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nofrt x0 nx0) (l_e_st_eq_landau_n_rt_nofrt y0 ny0) t) => l_e_st_eq_landau_n_rt_5112_t5 x0 nx0 y0 ny0 m t u))))))). Time Defined. (* constant 1714 *) Definition l_e_st_eq_landau_n_rt_satz112h : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_pf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl x y) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_picp (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y) (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl x y) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_satz112a x y))). Time Defined. (* constant 1715 *) Definition l_e_st_eq_landau_n_rt_satz112j : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_ts x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x y) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_tict (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y) (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x y) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_satz112b x y))). Time Defined. (* constant 1716 *) Definition l_e_st_eq_landau_n_rt_nt_natt : Type. exact (l_e_ot l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t)). Time Defined. (* constant 1717 *) Definition l_e_st_eq_landau_n_rt_nt_ntofrt : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), l_e_st_eq_landau_n_rt_nt_natt)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => l_e_out l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) x0 nx0)). Time Defined. (* constant 1718 *) Definition l_e_st_eq_landau_n_rt_nt_is : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_is l_e_st_eq_landau_n_rt_nt_natt xt yt)). Time Defined. (* constant 1719 *) Definition l_e_st_eq_landau_n_rt_nt_nis : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_not (l_e_st_eq_landau_n_rt_nt_is xt yt))). Time Defined. (* constant 1720 *) Definition l_e_st_eq_landau_n_rt_nt_all : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => l_all l_e_st_eq_landau_n_rt_nt_natt p). Time Defined. (* constant 1721 *) Definition l_e_st_eq_landau_n_rt_nt_some : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => l_some l_e_st_eq_landau_n_rt_nt_natt p). Time Defined. (* constant 1722 *) Definition l_e_st_eq_landau_n_rt_nt_one : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => l_e_one l_e_st_eq_landau_n_rt_nt_natt p). Time Defined. (* constant 1723 *) Definition l_e_st_eq_landau_n_rt_nt_in : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_esti l_e_st_eq_landau_n_rt_nt_natt xt st)). Time Defined. (* constant 1724 *) Definition l_e_st_eq_landau_n_rt_nt_rtofnt : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_rat). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_in l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) xt). Time Defined. (* constant 1725 *) Definition l_e_st_eq_landau_n_rt_nt_natrti : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_nt_rtofnt xt)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_inp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) xt). Time Defined. (* constant 1726 *) Definition l_e_st_eq_landau_n_rt_nt_isrtent : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofrt x0 nx0) (l_e_st_eq_landau_n_rt_nt_ntofrt y0 ny0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_isouti l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) x0 nx0 y0 ny0 i))))). Time Defined. (* constant 1727 *) Definition l_e_st_eq_landau_n_rt_nt_isrtint : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ny0:l_e_st_eq_landau_n_rt_natrt y0), (forall (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofrt x0 nx0) (l_e_st_eq_landau_n_rt_nt_ntofrt y0 ny0)), l_e_st_eq_landau_n_rt_is x0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ny0:l_e_st_eq_landau_n_rt_natrt y0) => (fun (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofrt x0 nx0) (l_e_st_eq_landau_n_rt_nt_ntofrt y0 ny0)) => l_e_isoute l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) x0 nx0 y0 ny0 i))))). Time Defined. (* constant 1728 *) Definition l_e_st_eq_landau_n_rt_nt_isntert : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_nt_is xt yt), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_nt_is xt yt) => l_e_isini l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) xt yt i))). Time Defined. (* constant 1729 *) Definition l_e_st_eq_landau_n_rt_nt_isntirt : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)), l_e_st_eq_landau_n_rt_nt_is xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) => l_e_isine l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) xt yt i))). Time Defined. (* constant 1730 *) Definition l_e_st_eq_landau_n_rt_nt_isrtnt1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (nx0:l_e_st_eq_landau_n_rt_natrt x0), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_nt_rtofnt (l_e_st_eq_landau_n_rt_nt_ntofrt x0 nx0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (nx0:l_e_st_eq_landau_n_rt_natrt x0) => l_e_isinout l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) x0 nx0)). Time Defined. (* constant 1731 *) Definition l_e_st_eq_landau_n_rt_nt_isntrt1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is xt (l_e_st_eq_landau_n_rt_nt_ntofrt (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_isoutin l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_natrt t) xt). Time Defined. (* constant 1732 *) Definition l_e_st_eq_landau_n_rt_nt_ntofn : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_nt_natt). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_ntofrt (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x)). Time Defined. (* constant 1733 *) Definition l_e_st_eq_landau_n_rt_nt_isnent : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofn x) (l_e_st_eq_landau_n_rt_nt_ntofn y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_rt_nt_isrtent (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x) (l_e_st_eq_landau_n_rt_rtofn y) (l_e_st_eq_landau_n_rt_natrti y) (l_e_st_eq_landau_n_rt_isnert x y i)))). Time Defined. (* constant 1734 *) Definition l_e_st_eq_landau_n_rt_nt_isnint : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofn x) (l_e_st_eq_landau_n_rt_nt_ntofn y)), l_e_st_eq_landau_n_is x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofn x) (l_e_st_eq_landau_n_rt_nt_ntofn y)) => l_e_st_eq_landau_n_rt_isnirt x y (l_e_st_eq_landau_n_rt_nt_isrtint (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x) (l_e_st_eq_landau_n_rt_rtofn y) (l_e_st_eq_landau_n_rt_natrti y) i)))). Time Defined. (* constant 1735 *) Definition l_e_st_eq_landau_n_rt_nt_nofnt : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_nat). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt)). Time Defined. (* constant 1736 *) Definition l_e_st_eq_landau_n_rt_nt_isnten : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_nt_is xt yt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_nt_is xt yt) => l_e_st_eq_landau_n_rt_isrten (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt) (l_e_st_eq_landau_n_rt_nt_isntert xt yt i)))). Time Defined. (* constant 1737 *) Definition l_e_st_eq_landau_n_rt_nt_isntin : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)), l_e_st_eq_landau_n_rt_nt_is xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_rt_nt_isntirt xt yt (l_e_st_eq_landau_n_rt_isrtin (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt) i)))). Time Defined. (* constant 1738 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t24 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_nt_rtofnt (l_e_st_eq_landau_n_rt_nt_ntofn x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_isrtnt1 (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x)). Time Defined. (* constant 1739 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t25 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_isrten (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x) (l_e_st_eq_landau_n_rt_nt_rtofnt (l_e_st_eq_landau_n_rt_nt_ntofn x)) (l_e_st_eq_landau_n_rt_nt_natrti (l_e_st_eq_landau_n_rt_nt_ntofn x)) (l_e_st_eq_landau_n_rt_nt_ii5_t24 x)). Time Defined. (* constant 1740 *) Definition l_e_st_eq_landau_n_rt_nt_isnnt1 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_natrti x)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn x)) (l_e_st_eq_landau_n_rt_isnrt1 x) (l_e_st_eq_landau_n_rt_nt_ii5_t25 x)). Time Defined. (* constant 1741 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t26 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_nt_nofnt xt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_isrtn1 (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt)). Time Defined. (* constant 1742 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t27 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofrt (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt)) (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_nofnt xt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_isrtent (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_rt_natrti (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_rt_nt_ii5_t26 xt)). Time Defined. (* constant 1743 *) Definition l_e_st_eq_landau_n_rt_nt_isntn1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is xt (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_nofnt xt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_tris l_e_st_eq_landau_n_rt_nt_natt xt (l_e_st_eq_landau_n_rt_nt_ntofrt (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt)) (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_rt_nt_isntrt1 xt) (l_e_st_eq_landau_n_rt_nt_ii5_t27 xt)). Time Defined. (* constant 1744 *) Definition l_e_st_eq_landau_n_rt_nt_isnnt2 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn x)) x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn x)) (l_e_st_eq_landau_n_rt_nt_isnnt1 x)). Time Defined. (* constant 1745 *) Definition l_e_st_eq_landau_n_rt_nt_isntn2 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_nofnt xt)) xt). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_symis l_e_st_eq_landau_n_rt_nt_natt xt (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_rt_nt_isntn1 xt)). Time Defined. (* constant 1746 *) Definition l_e_st_eq_landau_n_rt_nt_1t : l_e_st_eq_landau_n_rt_nt_natt. exact (l_e_st_eq_landau_n_rt_nt_ntofn l_e_st_eq_landau_n_1). Time Defined. (* constant 1747 *) Definition l_e_st_eq_landau_n_rt_nt_suct : (forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt). exact (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt x))). Time Defined. (* constant 1748 *) Definition l_e_st_eq_landau_n_rt_nt_5113_t1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) l_e_st_eq_landau_n_rt_nt_1t), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt xt)) l_e_st_eq_landau_n_1)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) l_e_st_eq_landau_n_rt_nt_1t) => l_e_st_eq_landau_n_rt_nt_isnint (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt xt)) l_e_st_eq_landau_n_1 i)). Time Defined. (* constant 1749 *) Definition l_e_st_eq_landau_n_rt_nt_satz113a : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_nis (l_e_st_eq_landau_n_rt_nt_suct xt) l_e_st_eq_landau_n_rt_nt_1t). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_imp_th3 (l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt xt)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ax3 (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (fun (t:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) l_e_st_eq_landau_n_rt_nt_1t) => l_e_st_eq_landau_n_rt_nt_5113_t1 xt t)). Time Defined. (* constant 1750 *) Definition l_e_st_eq_landau_n_rt_nt_5113_t2 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) (l_e_st_eq_landau_n_rt_nt_suct yt)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt yt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) (l_e_st_eq_landau_n_rt_nt_suct yt)) => l_e_st_eq_landau_n_rt_nt_isnint (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt yt)) i))). Time Defined. (* constant 1751 *) Definition l_e_st_eq_landau_n_rt_nt_satz113b : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) (l_e_st_eq_landau_n_rt_nt_suct yt)), l_e_st_eq_landau_n_rt_nt_is xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) (l_e_st_eq_landau_n_rt_nt_suct yt)) => l_e_st_eq_landau_n_rt_nt_isntin xt yt (l_e_st_eq_landau_n_ax4 (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_5113_t2 xt yt i))))). Time Defined. (* constant 1752 *) Definition l_e_st_eq_landau_n_rt_nt_cond1 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), Prop). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_in l_e_st_eq_landau_n_rt_nt_1t st). Time Defined. (* constant 1753 *) Definition l_e_st_eq_landau_n_rt_nt_cond2 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), Prop). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_all (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_imp (l_e_st_eq_landau_n_rt_nt_in x st) (l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_suct x) st))). Time Defined. (* constant 1754 *) Definition l_e_st_eq_landau_n_rt_nt_5113_prop1 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (x:l_e_st_eq_landau_n_nat), Prop)))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_ntofn x) st)))). Time Defined. (* constant 1755 *) Definition l_e_st_eq_landau_n_rt_nt_5113_t3 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 x), l_imp (l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_ntofn x) st) (l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_ntofn x)) st)))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 x) => c2 (l_e_st_eq_landau_n_rt_nt_ntofn x)))))). Time Defined. (* constant 1756 *) Definition l_e_st_eq_landau_n_rt_nt_5113_t4 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 x), l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_ntofn x)) st))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 x) => l_mp (l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_ntofn x) st) (l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_ntofn x)) st) p (l_e_st_eq_landau_n_rt_nt_5113_t3 st c1 c2 x p)))))). Time Defined. (* constant 1757 *) Definition l_e_st_eq_landau_n_rt_nt_5113_t5 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 x), l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 (l_e_st_eq_landau_n_suc x)))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 x) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_suc t)) st) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn x)) x (l_e_st_eq_landau_n_rt_nt_5113_t4 st c1 c2 x p) (l_e_st_eq_landau_n_rt_nt_isnnt2 x)))))). Time Defined. (* constant 1758 *) Definition l_e_st_eq_landau_n_rt_nt_5113_t6 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_in (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_nofnt xt)) st)))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 t) c1 (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_nt_5113_prop1 st c1 c2 t) => l_e_st_eq_landau_n_rt_nt_5113_t5 st c1 c2 t u)) (l_e_st_eq_landau_n_rt_nt_nofnt xt))))). Time Defined. (* constant 1759 *) Definition l_e_st_eq_landau_n_rt_nt_satz113c : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_in xt st)))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_isp l_e_st_eq_landau_n_rt_nt_natt (fun (t:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_in t st) (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_nofnt xt)) xt (l_e_st_eq_landau_n_rt_nt_5113_t6 st c1 c2 xt) (l_e_st_eq_landau_n_rt_nt_isntn2 xt))))). Time Defined. (* constant 1760 *) Definition l_e_st_eq_landau_n_rt_nt_ax3t : (forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_nis (l_e_st_eq_landau_n_rt_nt_suct x) l_e_st_eq_landau_n_rt_nt_1t). exact (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_satz113a x). Time Defined. (* constant 1761 *) Definition l_e_st_eq_landau_n_rt_nt_ax4t : (forall (x:l_e_st_eq_landau_n_rt_nt_natt), (forall (y:l_e_st_eq_landau_n_rt_nt_natt), (forall (u:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct x) (l_e_st_eq_landau_n_rt_nt_suct y)), l_e_st_eq_landau_n_rt_nt_is x y))). exact (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => (fun (y:l_e_st_eq_landau_n_rt_nt_natt) => (fun (u:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct x) (l_e_st_eq_landau_n_rt_nt_suct y)) => l_e_st_eq_landau_n_rt_nt_satz113b x y u))). Time Defined. (* constant 1762 *) Definition l_e_st_eq_landau_n_rt_nt_ax5t : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (u:l_e_st_eq_landau_n_rt_nt_cond1 s), (forall (v:l_e_st_eq_landau_n_rt_nt_cond2 s), (forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_in x s)))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (u:l_e_st_eq_landau_n_rt_nt_cond1 s) => (fun (v:l_e_st_eq_landau_n_rt_nt_cond2 s) => (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_satz113c s u v x)))). Time Defined. (* constant 1763 *) Definition l_e_st_eq_landau_n_rt_nt_51_t1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (n:l_e_st_eq_landau_n_rt_nt_nis xt yt), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (n:l_e_st_eq_landau_n_rt_nt_nis xt yt) => l_imp_th3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_nt_is xt yt) n (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_rt_nt_isntin xt yt t)))). Time Defined. (* constant 1764 *) Definition l_e_st_eq_landau_n_rt_nt_51_t2 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (n:l_e_st_eq_landau_n_rt_nt_nis xt yt), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt yt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (n:l_e_st_eq_landau_n_rt_nt_nis xt yt) => l_e_st_eq_landau_n_satz1 (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_51_t1 xt yt n)))). Time Defined. (* constant 1765 *) Definition l_e_st_eq_landau_n_rt_nt_satz1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (n:l_e_st_eq_landau_n_rt_nt_nis xt yt), l_e_st_eq_landau_n_rt_nt_nis (l_e_st_eq_landau_n_rt_nt_suct xt) (l_e_st_eq_landau_n_rt_nt_suct yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (n:l_e_st_eq_landau_n_rt_nt_nis xt yt) => l_imp_th3 (l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) (l_e_st_eq_landau_n_rt_nt_suct yt)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt yt))) (l_e_st_eq_landau_n_rt_nt_51_t2 xt yt n) (fun (t:l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_suct xt) (l_e_st_eq_landau_n_rt_nt_suct yt)) => l_e_st_eq_landau_n_rt_nt_isnint (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt xt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt yt)) t)))). Time Defined. (* constant 1766 *) Definition l_e_st_eq_landau_n_rt_nt_54_x : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_nat)))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_nofnt xt)))). Time Defined. (* constant 1767 *) Definition l_e_st_eq_landau_n_rt_nt_54_prop1t : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), Prop))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_e_st_eq_landau_n_rt_nt_all (fun (t:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_is (ft (l_e_st_eq_landau_n_rt_nt_suct t)) (l_e_st_eq_landau_n_rt_nt_suct (ft t)))))))). Time Defined. (* constant 1768 *) Definition l_e_st_eq_landau_n_rt_nt_54_prop2t : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), Prop))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_and (l_e_st_eq_landau_n_rt_nt_is (ft l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_rt_nt_suct xt)) (l_e_st_eq_landau_n_rt_nt_54_prop1t st c1 c2 xt ft)))))). Time Defined. (* constant 1769 *) Definition l_e_st_eq_landau_n_rt_nt_54_prop1 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), Prop))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_all (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_suc t)) (l_e_st_eq_landau_n_suc (f t)))))))). Time Defined. (* constant 1770 *) Definition l_e_st_eq_landau_n_rt_nt_54_prop2 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), Prop))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_and (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt))) (l_e_st_eq_landau_n_rt_nt_54_prop1 st c1 c2 xt f)))))). Time Defined. (* constant 1771 *) Definition l_e_st_eq_landau_n_rt_nt_54_g : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_nofnt (ft (l_e_st_eq_landau_n_rt_nt_ntofn t)))))))). Time Defined. (* constant 1772 *) Definition l_e_st_eq_landau_n_rt_nt_54_t1 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), l_e_st_eq_landau_n_rt_nt_is (ft l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_rt_nt_suct xt))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => l_ande1 (l_e_st_eq_landau_n_rt_nt_is (ft l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_rt_nt_suct xt)) (l_e_st_eq_landau_n_rt_nt_54_prop1t st c1 c2 xt ft) p)))))). Time Defined. (* constant 1773 *) Definition l_e_st_eq_landau_n_rt_nt_54_t2 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt)))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_suct xt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt)) (l_e_st_eq_landau_n_rt_nt_isnten (ft l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_rt_nt_suct xt) (l_e_st_eq_landau_n_rt_nt_54_t1 st c1 c2 xt ft p)) (l_e_st_eq_landau_n_rt_nt_isnnt2 (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt))))))))). Time Defined. (* constant 1774 *) Definition l_e_st_eq_landau_n_rt_nt_54_t3 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), l_e_st_eq_landau_n_rt_nt_54_prop1t st c1 c2 xt ft)))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => l_ande2 (l_e_st_eq_landau_n_rt_nt_is (ft l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_rt_nt_suct xt)) (l_e_st_eq_landau_n_rt_nt_54_prop1t st c1 c2 xt ft) p)))))). Time Defined. (* constant 1775 *) Definition l_e_st_eq_landau_n_rt_nt_54_ut : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_nt_natt))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_ntofn u))))))). Time Defined. (* constant 1776 *) Definition l_e_st_eq_landau_n_rt_nt_54_t4 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc u) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u)))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_suc u (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u)) (l_e_st_eq_landau_n_rt_nt_isnnt1 u)))))))). Time Defined. (* constant 1777 *) Definition l_e_st_eq_landau_n_rt_nt_54_t5 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_rt_nt_nofnt (ft (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u))))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft) (l_e_st_eq_landau_n_suc u) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u))) (l_e_st_eq_landau_n_rt_nt_54_t4 st c1 c2 xt ft p u)))))))). Time Defined. (* constant 1778 *) Definition l_e_st_eq_landau_n_rt_nt_54_t6 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_nt_is (ft (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u))) (l_e_st_eq_landau_n_rt_nt_suct (ft (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u)))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_54_t3 st c1 c2 xt ft p (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u)))))))). Time Defined. (* constant 1779 *) Definition l_e_st_eq_landau_n_rt_nt_54_t7 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt (ft (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u)))) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_suct (ft (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u))))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_isnten (ft (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u))) (l_e_st_eq_landau_n_rt_nt_suct (ft (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u))) (l_e_st_eq_landau_n_rt_nt_54_t6 st c1 c2 xt ft p u)))))))). Time Defined. (* constant 1780 *) Definition l_e_st_eq_landau_n_rt_nt_54_t8 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_suct (ft (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u)))) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft u))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_isnnt2 (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft u))))))))). Time Defined. (* constant 1781 *) Definition l_e_st_eq_landau_n_rt_nt_54_t9 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft u))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_rt_nt_nofnt (ft (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u)))) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_suct (ft (l_e_st_eq_landau_n_rt_nt_54_ut st c1 c2 xt ft p u)))) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft u)) (l_e_st_eq_landau_n_rt_nt_54_t5 st c1 c2 xt ft p u) (l_e_st_eq_landau_n_rt_nt_54_t7 st c1 c2 xt ft p u) (l_e_st_eq_landau_n_rt_nt_54_t8 st c1 c2 xt ft p u)))))))). Time Defined. (* constant 1782 *) Definition l_e_st_eq_landau_n_rt_nt_54_t10 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), l_e_st_eq_landau_n_rt_nt_54_prop1 st c1 c2 xt (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_nt_54_t9 st c1 c2 xt ft p u))))))). Time Defined. (* constant 1783 *) Definition l_e_st_eq_landau_n_rt_nt_54_t11 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft), l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ft:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt ft) => l_andi (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt))) (l_e_st_eq_landau_n_rt_nt_54_prop1 st c1 c2 xt (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt ft)) (l_e_st_eq_landau_n_rt_nt_54_t2 st c1 c2 xt ft p) (l_e_st_eq_landau_n_rt_nt_54_t10 st c1 c2 xt ft p))))))). Time Defined. (* constant 1784 *) Definition l_e_st_eq_landau_n_rt_nt_54_t12 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a), (forall (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b), l_e_amone (forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (u:(forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt u))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a) => (fun (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b) => l_e_onee1 (forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (u:(forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt u) (l_e_st_eq_landau_n_satz4 (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt)))))))))). Time Defined. (* constant 1785 *) Definition l_e_st_eq_landau_n_rt_nt_54_t13 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a), (forall (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b), l_e_is (forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt a) (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt b))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a) => (fun (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b) => l_e_st_eq_landau_n_rt_nt_54_t12 st c1 c2 xt a b pa pb (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt a) (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt b) (l_e_st_eq_landau_n_rt_nt_54_t11 st c1 c2 xt a pa) (l_e_st_eq_landau_n_rt_nt_54_t11 st c1 c2 xt b pb))))))))). Time Defined. (* constant 1786 *) Definition l_e_st_eq_landau_n_rt_nt_54_y : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a), (forall (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_nat))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a) => (fun (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_nofnt yt))))))))). Time Defined. (* constant 1787 *) Definition l_e_st_eq_landau_n_rt_nt_54_t14 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a), (forall (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt (a (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt)))) (l_e_st_eq_landau_n_rt_nt_nofnt (b (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt))))))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a) => (fun (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_fise l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt a) (l_e_st_eq_landau_n_rt_nt_54_g st c1 c2 xt b) (l_e_st_eq_landau_n_rt_nt_54_t13 st c1 c2 xt a b pa pb) (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt)))))))))). Time Defined. (* constant 1788 *) Definition l_e_st_eq_landau_n_rt_nt_54_t15 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a), (forall (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is (a (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt))) (b (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt)))))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a) => (fun (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_isntin (a (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt))) (b (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt))) (l_e_st_eq_landau_n_rt_nt_54_t14 st c1 c2 xt a b pa pb yt)))))))))). Time Defined. (* constant 1789 *) Definition l_e_st_eq_landau_n_rt_nt_54_t16 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a), (forall (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is (a yt) (b yt)))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a) => (fun (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_tr3is l_e_st_eq_landau_n_rt_nt_natt (a yt) (a (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt))) (b (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt))) (b yt) (l_e_isf l_e_st_eq_landau_n_rt_nt_natt l_e_st_eq_landau_n_rt_nt_natt a yt (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt)) (l_e_st_eq_landau_n_rt_nt_isntn1 yt)) (l_e_st_eq_landau_n_rt_nt_54_t15 st c1 c2 xt a b pa pb yt) (l_e_isf l_e_st_eq_landau_n_rt_nt_natt l_e_st_eq_landau_n_rt_nt_natt b (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_54_y st c1 c2 xt a b pa pb yt)) yt (l_e_st_eq_landau_n_rt_nt_isntn2 yt))))))))))). Time Defined. (* constant 1790 *) Definition l_e_st_eq_landau_n_rt_nt_54_t17 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)), (forall (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a), (forall (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b), l_e_is (forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt) a b)))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (a:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (b:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (pa:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt a) => (fun (pb:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt b) => l_e_fisi l_e_st_eq_landau_n_rt_nt_natt l_e_st_eq_landau_n_rt_nt_natt a b (fun (t:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_54_t16 st c1 c2 xt a b pa pb t))))))))). Time Defined. (* constant 1791 *) Definition l_e_st_eq_landau_n_rt_nt_54_t18 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_amone (forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt) (fun (u:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt u))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (u:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (v:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => (fun (w:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt u) => (fun (z:l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt v) => l_e_st_eq_landau_n_rt_nt_54_t17 st c1 c2 xt u v w z)))))))). Time Defined. (* constant 1792 *) Definition l_e_st_eq_landau_n_rt_nt_54_t19 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_some (forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (u:(forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt u))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_onee2 (forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (u:(forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt u) (l_e_st_eq_landau_n_satz4 (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt)))))). Time Defined. (* constant 1793 *) Definition l_e_st_eq_landau_n_rt_nt_54_gt : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (x:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (t:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_ntofn (f (l_e_st_eq_landau_n_rt_nt_nofnt t)))))))). Time Defined. (* constant 1794 *) Definition l_e_st_eq_landau_n_rt_nt_54_t20 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt)))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => l_ande1 (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt))) (l_e_st_eq_landau_n_rt_nt_54_prop1 st c1 c2 xt f) p)))))). Time Defined. (* constant 1795 *) Definition l_e_st_eq_landau_n_rt_nt_54_t21 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_rt_nt_nofnt l_e_st_eq_landau_n_rt_nt_1t)) (f l_e_st_eq_landau_n_1))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat f (l_e_st_eq_landau_n_rt_nt_nofnt l_e_st_eq_landau_n_rt_nt_1t) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_nt_isnnt2 l_e_st_eq_landau_n_1))))))). Time Defined. (* constant 1796 *) Definition l_e_st_eq_landau_n_rt_nt_54_t22 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_rt_nt_suct xt))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => l_e_st_eq_landau_n_rt_nt_isnent (f (l_e_st_eq_landau_n_rt_nt_nofnt l_e_st_eq_landau_n_rt_nt_1t)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt)) (l_e_tris l_e_st_eq_landau_n_nat (f (l_e_st_eq_landau_n_rt_nt_nofnt l_e_st_eq_landau_n_rt_nt_1t)) (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt)) (l_e_st_eq_landau_n_rt_nt_54_t21 st c1 c2 xt f p) (l_e_st_eq_landau_n_rt_nt_54_t20 st c1 c2 xt f p)))))))). Time Defined. (* constant 1797 *) Definition l_e_st_eq_landau_n_rt_nt_54_t23 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), l_e_st_eq_landau_n_rt_nt_54_prop1 st c1 c2 xt f)))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => l_ande2 (l_e_st_eq_landau_n_is (f l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_x st c1 c2 xt))) (l_e_st_eq_landau_n_rt_nt_54_prop1 st c1 c2 xt f) p)))))). Time Defined. (* constant 1798 *) Definition l_e_st_eq_landau_n_rt_nt_54_z : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_nat))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_nofnt zt))))))). Time Defined. (* constant 1799 *) Definition l_e_st_eq_landau_n_rt_nt_54_t24 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_suct zt))) (f (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt)))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat f (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_suct zt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt)) (l_e_st_eq_landau_n_rt_nt_isnnt2 (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt)))))))))). Time Defined. (* constant 1800 *) Definition l_e_st_eq_landau_n_rt_nt_54_t25 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (f (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt))) (l_e_st_eq_landau_n_suc (f (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt)))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_54_t23 st c1 c2 xt f p (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt)))))))). Time Defined. (* constant 1801 *) Definition l_e_st_eq_landau_n_rt_nt_54_t26 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (f (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt))) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f zt)))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_suc (f (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f zt)) (l_e_st_eq_landau_n_rt_nt_isnnt1 (f (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt)))))))))). Time Defined. (* constant 1802 *) Definition l_e_st_eq_landau_n_rt_nt_54_t27 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f (l_e_st_eq_landau_n_rt_nt_suct zt)) (l_e_st_eq_landau_n_rt_nt_suct (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f zt))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_isnent (f (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_suct zt))) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f zt))) (l_e_tr3is l_e_st_eq_landau_n_nat (f (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_suct zt))) (f (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt))) (l_e_st_eq_landau_n_suc (f (l_e_st_eq_landau_n_rt_nt_54_z st c1 c2 xt f p zt))) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f zt))) (l_e_st_eq_landau_n_rt_nt_54_t24 st c1 c2 xt f p zt) (l_e_st_eq_landau_n_rt_nt_54_t25 st c1 c2 xt f p zt) (l_e_st_eq_landau_n_rt_nt_54_t26 st c1 c2 xt f p zt))))))))). Time Defined. (* constant 1803 *) Definition l_e_st_eq_landau_n_rt_nt_54_t28 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), l_e_st_eq_landau_n_rt_nt_54_prop1t st c1 c2 xt (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_54_t27 st c1 c2 xt f p u))))))). Time Defined. (* constant 1804 *) Definition l_e_st_eq_landau_n_rt_nt_54_t29 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => l_andi (l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_rt_nt_suct xt)) (l_e_st_eq_landau_n_rt_nt_54_prop1t st c1 c2 xt (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f)) (l_e_st_eq_landau_n_rt_nt_54_t22 st c1 c2 xt f p) (l_e_st_eq_landau_n_rt_nt_54_t28 st c1 c2 xt f p))))))). Time Defined. (* constant 1805 *) Definition l_e_st_eq_landau_n_rt_nt_54_t30 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)), (forall (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f), l_some (forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt) (fun (u:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt u))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (f:(forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (p:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt f) => l_somei (forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt) (fun (u:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt u) (l_e_st_eq_landau_n_rt_nt_54_gt st c1 c2 xt f) (l_e_st_eq_landau_n_rt_nt_54_t29 st c1 c2 xt f p))))))). Time Defined. (* constant 1806 *) Definition l_e_st_eq_landau_n_rt_nt_54_t31 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_some (forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt) (fun (u:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt u))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_someapp (forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat) (fun (u:(forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt u) (l_e_st_eq_landau_n_rt_nt_54_t19 st c1 c2 xt) (l_some (forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt) (fun (u:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt u)) (fun (u:(forall (t:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)) => (fun (v:l_e_st_eq_landau_n_rt_nt_54_prop2 st c1 c2 xt u) => l_e_st_eq_landau_n_rt_nt_54_t30 st c1 c2 xt u v)))))). Time Defined. (* constant 1807 *) Definition l_e_st_eq_landau_n_rt_nt_satz4 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), l_e_one (forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt) (fun (u:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_and (l_e_st_eq_landau_n_rt_nt_is (u l_e_st_eq_landau_n_rt_nt_1t) (l_e_st_eq_landau_n_rt_nt_suct xt)) (l_e_st_eq_landau_n_rt_nt_all (fun (v:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_is (u (l_e_st_eq_landau_n_rt_nt_suct v)) (l_e_st_eq_landau_n_rt_nt_suct (u v))))))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_onei (forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt) (fun (u:(forall (t:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)) => l_e_st_eq_landau_n_rt_nt_54_prop2t st c1 c2 xt u) (l_e_st_eq_landau_n_rt_nt_54_t18 st c1 c2 xt) (l_e_st_eq_landau_n_rt_nt_54_t31 st c1 c2 xt))))). Time Defined. (* constant 1808 *) Definition l_e_st_eq_landau_n_rt_nt_pl : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_natt)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_ntofrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_rt_satz112d (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt)))). Time Defined. (* constant 1809 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t28 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_inf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_class (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_satz112c (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt))). Time Defined. (* constant 1810 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t29 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_nt_ii5_t28 xt yt) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_refeq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) l_e_st_eq_landau_n_1)))). Time Defined. (* constant 1811 *) Definition l_e_st_eq_landau_n_rt_nt_isplnt : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_pl xt yt) (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_isrtent (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_rt_satz112d (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))) (l_e_st_eq_landau_n_rt_natrti (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))) (l_e_st_eq_landau_n_rt_nt_ii5_t29 xt yt))). Time Defined. (* constant 1812 *) Definition l_e_st_eq_landau_n_rt_nt_isntpl : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))) (l_e_st_eq_landau_n_rt_nt_pl xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_symis l_e_st_eq_landau_n_rt_nt_natt (l_e_st_eq_landau_n_rt_nt_pl xt yt) (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))) (l_e_st_eq_landau_n_rt_nt_isplnt xt yt))). Time Defined. (* constant 1813 *) Definition l_e_st_eq_landau_n_rt_nt_ispln : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)))) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_rt_nt_isnnt1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))) (l_e_st_eq_landau_n_rt_nt_isnten (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))) (l_e_st_eq_landau_n_rt_nt_pl xt yt) (l_e_st_eq_landau_n_rt_nt_isntpl xt yt)))). Time Defined. (* constant 1814 *) Definition l_e_st_eq_landau_n_rt_nt_isnpl : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_rt_nt_ispln xt yt))). Time Defined. (* constant 1815 *) Definition l_e_st_eq_landau_n_rt_nt_55_t1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt zt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt zt) (l_e_st_eq_landau_n_rt_nt_isnpl xt yt)))). Time Defined. (* constant 1816 *) Definition l_e_st_eq_landau_n_rt_nt_55_t2 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt)))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt)) (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_ispln yt zt)))). Time Defined. (* constant 1817 *) Definition l_e_st_eq_landau_n_rt_nt_55_t3 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt)))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt))) (l_e_st_eq_landau_n_rt_nt_55_t1 xt yt zt) (l_e_st_eq_landau_n_satz5 (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_rt_nt_55_t2 xt yt zt)))). Time Defined. (* constant 1818 *) Definition l_e_st_eq_landau_n_rt_nt_satz5 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), l_e_st_eq_landau_n_rt_nt_is (l_e_st_eq_landau_n_rt_nt_pl (l_e_st_eq_landau_n_rt_nt_pl xt yt) zt) (l_e_st_eq_landau_n_rt_nt_pl xt (l_e_st_eq_landau_n_rt_nt_pl yt zt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_tr3is l_e_st_eq_landau_n_rt_nt_natt (l_e_st_eq_landau_n_rt_nt_pl (l_e_st_eq_landau_n_rt_nt_pl xt yt) zt) (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt zt))) (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt)))) (l_e_st_eq_landau_n_rt_nt_pl xt (l_e_st_eq_landau_n_rt_nt_pl yt zt)) (l_e_st_eq_landau_n_rt_nt_isplnt (l_e_st_eq_landau_n_rt_nt_pl xt yt) zt) (l_e_st_eq_landau_n_rt_nt_isnent (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt yt)) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt))) (l_e_st_eq_landau_n_rt_nt_55_t3 xt yt zt)) (l_e_st_eq_landau_n_rt_nt_isntpl xt (l_e_st_eq_landau_n_rt_nt_pl yt zt))))). Time Defined. (* constant 1819 *) Definition l_e_st_eq_landau_n_rt_nt_diffprop : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), Prop))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_is xt (l_e_st_eq_landau_n_rt_nt_pl yt zt)))). Time Defined. (* constant 1820 *) Definition l_e_st_eq_landau_n_rt_nt_diffprope : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), (forall (d:l_e_st_eq_landau_n_rt_nt_diffprop xt yt zt), l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (d:l_e_st_eq_landau_n_rt_nt_diffprop xt yt zt) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_rt_nt_isnten xt (l_e_st_eq_landau_n_rt_nt_pl yt zt) d) (l_e_st_eq_landau_n_rt_nt_isnpl yt zt))))). Time Defined. (* constant 1821 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t30 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), (forall (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt)))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt zt)) d (l_e_st_eq_landau_n_rt_nt_ispln yt zt))))). Time Defined. (* constant 1822 *) Definition l_e_st_eq_landau_n_rt_nt_diffpropi : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), (forall (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)), l_e_st_eq_landau_n_rt_nt_diffprop xt yt zt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) => l_e_st_eq_landau_n_rt_nt_isntin xt (l_e_st_eq_landau_n_rt_nt_pl yt zt) (l_e_st_eq_landau_n_rt_nt_ii5_t30 xt yt zt d))))). Time Defined. (* constant 1823 *) Definition l_e_st_eq_landau_n_rt_nt_59_it : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_is xt yt)). Time Defined. (* constant 1824 *) Definition l_e_st_eq_landau_n_rt_nt_59_iit : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_some (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_diffprop xt yt u))). Time Defined. (* constant 1825 *) Definition l_e_st_eq_landau_n_rt_nt_59_iiit : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_some (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_diffprop yt xt u))). Time Defined. (* constant 1826 *) Definition l_e_st_eq_landau_n_rt_nt_59_i : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))). Time Defined. (* constant 1827 *) Definition l_e_st_eq_landau_n_rt_nt_59_ii : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) u))). Time Defined. (* constant 1828 *) Definition l_e_st_eq_landau_n_rt_nt_59_iii : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt xt) u))). Time Defined. (* constant 1829 *) Definition l_e_st_eq_landau_n_rt_nt_59_t1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (one:l_e_st_eq_landau_n_rt_nt_59_i xt yt), l_or3 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (one:l_e_st_eq_landau_n_rt_nt_59_i xt yt) => l_or3i1 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (l_e_st_eq_landau_n_rt_nt_isntin xt yt one)))). Time Defined. (* constant 1830 *) Definition l_e_st_eq_landau_n_rt_nt_59_t2 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (two:l_e_st_eq_landau_n_rt_nt_59_ii xt yt), (forall (v:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) v), l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn v))))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (two:l_e_st_eq_landau_n_rt_nt_59_ii xt yt) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) v) => l_e_isp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) u) v (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn v)) d (l_e_st_eq_landau_n_rt_nt_isnnt1 v)))))). Time Defined. (* constant 1831 *) Definition l_e_st_eq_landau_n_rt_nt_59_t3 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (two:l_e_st_eq_landau_n_rt_nt_59_ii xt yt), (forall (v:l_e_st_eq_landau_n_nat), (forall (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) v), l_e_st_eq_landau_n_rt_nt_59_iit xt yt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (two:l_e_st_eq_landau_n_rt_nt_59_ii xt yt) => (fun (v:l_e_st_eq_landau_n_nat) => (fun (d:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) v) => l_somei l_e_st_eq_landau_n_rt_nt_natt (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_diffprop xt yt u) (l_e_st_eq_landau_n_rt_nt_ntofn v) (l_e_st_eq_landau_n_rt_nt_diffpropi xt yt (l_e_st_eq_landau_n_rt_nt_ntofn v) (l_e_st_eq_landau_n_rt_nt_59_t2 xt yt two v d))))))). Time Defined. (* constant 1832 *) Definition l_e_st_eq_landau_n_rt_nt_59_t4 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (two:l_e_st_eq_landau_n_rt_nt_59_ii xt yt), l_e_st_eq_landau_n_rt_nt_59_iit xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (two:l_e_st_eq_landau_n_rt_nt_59_ii xt yt) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) u) two (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) u) => l_e_st_eq_landau_n_rt_nt_59_t3 xt yt two u v))))). Time Defined. (* constant 1833 *) Definition l_e_st_eq_landau_n_rt_nt_59_t5 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (two:l_e_st_eq_landau_n_rt_nt_59_ii xt yt), l_or3 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (two:l_e_st_eq_landau_n_rt_nt_59_ii xt yt) => l_or3i2 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (l_e_st_eq_landau_n_rt_nt_59_t4 xt yt two)))). Time Defined. (* constant 1834 *) Definition l_e_st_eq_landau_n_rt_nt_59_t6 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (three:l_e_st_eq_landau_n_rt_nt_59_iii xt yt), l_or3 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (three:l_e_st_eq_landau_n_rt_nt_59_iii xt yt) => l_or3i3 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (l_e_st_eq_landau_n_rt_nt_59_t4 yt xt three)))). Time Defined. (* constant 1835 *) Definition l_e_st_eq_landau_n_rt_nt_59_t7 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_or3 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_or3app (l_e_st_eq_landau_n_rt_nt_59_i xt yt) (l_e_st_eq_landau_n_rt_nt_59_ii xt yt) (l_e_st_eq_landau_n_rt_nt_59_iii xt yt) (l_or3 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt)) (l_e_st_eq_landau_n_satz9a (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (fun (u:l_e_st_eq_landau_n_rt_nt_59_i xt yt) => l_e_st_eq_landau_n_rt_nt_59_t1 xt yt u) (fun (u:l_e_st_eq_landau_n_rt_nt_59_ii xt yt) => l_e_st_eq_landau_n_rt_nt_59_t5 xt yt u) (fun (u:l_e_st_eq_landau_n_rt_nt_59_iii xt yt) => l_e_st_eq_landau_n_rt_nt_59_t6 xt yt u))). Time Defined. (* constant 1836 *) Definition l_e_st_eq_landau_n_rt_nt_59_t8 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (onet:l_e_st_eq_landau_n_rt_nt_59_it xt yt), l_e_st_eq_landau_n_rt_nt_59_i xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (onet:l_e_st_eq_landau_n_rt_nt_59_it xt yt) => l_e_st_eq_landau_n_rt_nt_isnten xt yt onet))). Time Defined. (* constant 1837 *) Definition l_e_st_eq_landau_n_rt_nt_59_t9 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (twot:l_e_st_eq_landau_n_rt_nt_59_iit xt yt), (forall (vt:l_e_st_eq_landau_n_rt_nt_natt), (forall (d:l_e_st_eq_landau_n_rt_nt_diffprop xt yt vt), l_e_st_eq_landau_n_rt_nt_59_ii xt yt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (twot:l_e_st_eq_landau_n_rt_nt_59_iit xt yt) => (fun (vt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (d:l_e_st_eq_landau_n_rt_nt_diffprop xt yt vt) => l_somei l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) u) (l_e_st_eq_landau_n_rt_nt_nofnt vt) (l_e_st_eq_landau_n_rt_nt_diffprope xt yt vt d)))))). Time Defined. (* constant 1838 *) Definition l_e_st_eq_landau_n_rt_nt_59_t10 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (twot:l_e_st_eq_landau_n_rt_nt_59_iit xt yt), l_e_st_eq_landau_n_rt_nt_59_ii xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (twot:l_e_st_eq_landau_n_rt_nt_59_iit xt yt) => l_someapp l_e_st_eq_landau_n_rt_nt_natt (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_diffprop xt yt u) twot (l_e_st_eq_landau_n_rt_nt_59_ii xt yt) (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => (fun (v:l_e_st_eq_landau_n_rt_nt_diffprop xt yt u) => l_e_st_eq_landau_n_rt_nt_59_t9 xt yt twot u v))))). Time Defined. (* constant 1839 *) Definition l_e_st_eq_landau_n_rt_nt_59_t11 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (threet:l_e_st_eq_landau_n_rt_nt_59_iiit xt yt), l_e_st_eq_landau_n_rt_nt_59_iii xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (threet:l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) => l_e_st_eq_landau_n_rt_nt_59_t10 yt xt threet))). Time Defined. (* constant 1840 *) Definition l_e_st_eq_landau_n_rt_nt_59_t12 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_ec3 (l_e_st_eq_landau_n_rt_nt_59_i xt yt) (l_e_st_eq_landau_n_rt_nt_59_ii xt yt) (l_e_st_eq_landau_n_rt_nt_59_iii xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_satz9b (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt))). Time Defined. (* constant 1841 *) Definition l_e_st_eq_landau_n_rt_nt_59_t13 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (onet:l_e_st_eq_landau_n_rt_nt_59_it xt yt), l_not (l_e_st_eq_landau_n_rt_nt_59_ii xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (onet:l_e_st_eq_landau_n_rt_nt_59_it xt yt) => l_ec3e12 (l_e_st_eq_landau_n_rt_nt_59_i xt yt) (l_e_st_eq_landau_n_rt_nt_59_ii xt yt) (l_e_st_eq_landau_n_rt_nt_59_iii xt yt) (l_e_st_eq_landau_n_rt_nt_59_t12 xt yt) (l_e_st_eq_landau_n_rt_nt_59_t8 xt yt onet)))). Time Defined. (* constant 1842 *) Definition l_e_st_eq_landau_n_rt_nt_59_t14 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (onet:l_e_st_eq_landau_n_rt_nt_59_it xt yt), l_not (l_e_st_eq_landau_n_rt_nt_59_iit xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (onet:l_e_st_eq_landau_n_rt_nt_59_it xt yt) => l_imp_th3 (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_ii xt yt) (l_e_st_eq_landau_n_rt_nt_59_t13 xt yt onet) (fun (x:l_e_st_eq_landau_n_rt_nt_59_iit xt yt) => l_e_st_eq_landau_n_rt_nt_59_t10 xt yt x)))). Time Defined. (* constant 1843 *) Definition l_e_st_eq_landau_n_rt_nt_59_t15 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_ec (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_ec_th1 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (fun (x:l_e_st_eq_landau_n_rt_nt_59_it xt yt) => l_e_st_eq_landau_n_rt_nt_59_t14 xt yt x))). Time Defined. (* constant 1844 *) Definition l_e_st_eq_landau_n_rt_nt_59_t16 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (twot:l_e_st_eq_landau_n_rt_nt_59_iit xt yt), l_not (l_e_st_eq_landau_n_rt_nt_59_iii xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (twot:l_e_st_eq_landau_n_rt_nt_59_iit xt yt) => l_ec3e23 (l_e_st_eq_landau_n_rt_nt_59_i xt yt) (l_e_st_eq_landau_n_rt_nt_59_ii xt yt) (l_e_st_eq_landau_n_rt_nt_59_iii xt yt) (l_e_st_eq_landau_n_rt_nt_59_t12 xt yt) (l_e_st_eq_landau_n_rt_nt_59_t10 xt yt twot)))). Time Defined. (* constant 1845 *) Definition l_e_st_eq_landau_n_rt_nt_59_t17 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (twot:l_e_st_eq_landau_n_rt_nt_59_iit xt yt), l_not (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (twot:l_e_st_eq_landau_n_rt_nt_59_iit xt yt) => l_imp_th3 (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iii xt yt) (l_e_st_eq_landau_n_rt_nt_59_t16 xt yt twot) (fun (x:l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) => l_e_st_eq_landau_n_rt_nt_59_t11 xt yt x)))). Time Defined. (* constant 1846 *) Definition l_e_st_eq_landau_n_rt_nt_59_t18 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_ec (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_ec_th1 (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (fun (x:l_e_st_eq_landau_n_rt_nt_59_iit xt yt) => l_e_st_eq_landau_n_rt_nt_59_t17 xt yt x))). Time Defined. (* constant 1847 *) Definition l_e_st_eq_landau_n_rt_nt_59_t19 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (threet:l_e_st_eq_landau_n_rt_nt_59_iiit xt yt), l_not (l_e_st_eq_landau_n_rt_nt_59_i xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (threet:l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) => l_ec3e31 (l_e_st_eq_landau_n_rt_nt_59_i xt yt) (l_e_st_eq_landau_n_rt_nt_59_ii xt yt) (l_e_st_eq_landau_n_rt_nt_59_iii xt yt) (l_e_st_eq_landau_n_rt_nt_59_t12 xt yt) (l_e_st_eq_landau_n_rt_nt_59_t11 xt yt threet)))). Time Defined. (* constant 1848 *) Definition l_e_st_eq_landau_n_rt_nt_59_t20 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (threet:l_e_st_eq_landau_n_rt_nt_59_iiit xt yt), l_not (l_e_st_eq_landau_n_rt_nt_59_it xt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (threet:l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) => l_imp_th3 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_i xt yt) (l_e_st_eq_landau_n_rt_nt_59_t19 xt yt threet) (fun (x:l_e_st_eq_landau_n_rt_nt_59_it xt yt) => l_e_st_eq_landau_n_rt_nt_59_t8 xt yt x)))). Time Defined. (* constant 1849 *) Definition l_e_st_eq_landau_n_rt_nt_59_t21 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_ec (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (l_e_st_eq_landau_n_rt_nt_59_it xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_ec_th1 (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (fun (x:l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) => l_e_st_eq_landau_n_rt_nt_59_t20 xt yt x))). Time Defined. (* constant 1850 *) Definition l_e_st_eq_landau_n_rt_nt_59_t22 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_ec3 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_ec3_th6 (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (l_e_st_eq_landau_n_rt_nt_59_t15 xt yt) (l_e_st_eq_landau_n_rt_nt_59_t18 xt yt) (l_e_st_eq_landau_n_rt_nt_59_t21 xt yt))). Time Defined. (* constant 1851 *) Definition l_e_st_eq_landau_n_rt_nt_satz9 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), l_orec3 (l_e_st_eq_landau_n_rt_nt_is xt yt) (l_e_st_eq_landau_n_rt_nt_some (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_is xt (l_e_st_eq_landau_n_rt_nt_pl yt u))) (l_e_st_eq_landau_n_rt_nt_some (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_is yt (l_e_st_eq_landau_n_rt_nt_pl xt u))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_orec3i (l_e_st_eq_landau_n_rt_nt_59_it xt yt) (l_e_st_eq_landau_n_rt_nt_59_iit xt yt) (l_e_st_eq_landau_n_rt_nt_59_iiit xt yt) (l_e_st_eq_landau_n_rt_nt_59_t7 xt yt) (l_e_st_eq_landau_n_rt_nt_59_t22 xt yt))). Time Defined. (* constant 1852 *) Definition l_e_st_eq_landau_n_rt_nt_more : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt))). Time Defined. (* constant 1853 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t31 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_rt_nt_more xt yt), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt xt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt yt) l_e_st_eq_landau_n_1)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_rt_nt_more xt yt) => l_e_st_eq_landau_n_rt_moree (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt xt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt yt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt)) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt)) m))). Time Defined. (* constant 1854 *) Definition l_e_st_eq_landau_n_rt_nt_moree : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_rt_nt_more xt yt), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_rt_nt_more xt yt) => l_e_st_eq_landau_n_satz111a (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_ii5_t31 xt yt m)))). Time Defined. (* constant 1855 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t32 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt xt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt yt) l_e_st_eq_landau_n_1)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_satz111d (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) m))). Time Defined. (* constant 1856 *) Definition l_e_st_eq_landau_n_rt_nt_morei : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)), l_e_st_eq_landau_n_rt_nt_more xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt xt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt yt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt)) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt)) (l_e_st_eq_landau_n_rt_nt_ii5_t32 xt yt m)))). Time Defined. (* constant 1857 *) Definition l_e_st_eq_landau_n_rt_nt_less : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt))). Time Defined. (* constant 1858 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t33 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (l:l_e_st_eq_landau_n_rt_nt_less xt yt), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt xt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt yt) l_e_st_eq_landau_n_1)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (l:l_e_st_eq_landau_n_rt_nt_less xt yt) => l_e_st_eq_landau_n_rt_lesse (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt xt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt yt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt)) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt)) l))). Time Defined. (* constant 1859 *) Definition l_e_st_eq_landau_n_rt_nt_lesse : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (l:l_e_st_eq_landau_n_rt_nt_less xt yt), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (l:l_e_st_eq_landau_n_rt_nt_less xt yt) => l_e_st_eq_landau_n_satz111c (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_ii5_t33 xt yt l)))). Time Defined. (* constant 1860 *) Definition l_e_st_eq_landau_n_rt_nt_ii5_t34 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt xt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt yt) l_e_st_eq_landau_n_1)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_satz111f (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) l))). Time Defined. (* constant 1861 *) Definition l_e_st_eq_landau_n_rt_nt_lessi : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)), l_e_st_eq_landau_n_rt_nt_less xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_rt_lessi (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt xt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_nt_nofnt yt) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt)) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt)) (l_e_st_eq_landau_n_rt_nt_ii5_t34 xt yt l)))). Time Defined. (* constant 1862 *) Definition l_e_st_eq_landau_n_rt_nt_moreis : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt))). Time Defined. (* constant 1863 *) Definition l_e_st_eq_landau_n_rt_nt_moreise : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_rt_nt_moreis xt yt), l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_rt_nt_moreis xt yt) => l_or_th9 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) m (fun (u:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) => l_e_st_eq_landau_n_rt_nt_moree xt yt u) (fun (u:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) => l_e_st_eq_landau_n_rt_isrten (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt) u)))). Time Defined. (* constant 1864 *) Definition l_e_st_eq_landau_n_rt_nt_moreisi : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)), l_e_st_eq_landau_n_rt_nt_moreis xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_or_th9 (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) m (fun (u:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_rt_nt_morei xt yt u) (fun (u:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_rt_isrtin (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt) u)))). Time Defined. (* constant 1865 *) Definition l_e_st_eq_landau_n_rt_nt_lessis : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_lessis (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt))). Time Defined. (* constant 1866 *) Definition l_e_st_eq_landau_n_rt_nt_lessise : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (l:l_e_st_eq_landau_n_rt_nt_lessis xt yt), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (l:l_e_st_eq_landau_n_rt_nt_lessis xt yt) => l_or_th9 (l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) l (fun (u:l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) => l_e_st_eq_landau_n_rt_nt_lesse xt yt u) (fun (u:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) => l_e_st_eq_landau_n_rt_isrten (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt) u)))). Time Defined. (* constant 1867 *) Definition l_e_st_eq_landau_n_rt_nt_lessisi : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)), l_e_st_eq_landau_n_rt_nt_lessis xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_or_th9 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) (l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt)) l (fun (u:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_rt_nt_lessi xt yt u) (fun (u:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt)) => l_e_st_eq_landau_n_rt_isrtin (l_e_st_eq_landau_n_rt_nt_rtofnt xt) (l_e_st_eq_landau_n_rt_nt_natrti xt) (l_e_st_eq_landau_n_rt_nt_rtofnt yt) (l_e_st_eq_landau_n_rt_nt_natrti yt) u)))). Time Defined. (* constant 1868 *) Definition l_e_st_eq_landau_n_rt_nt_515_t1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), (forall (l:l_e_st_eq_landau_n_rt_nt_less xt yt), (forall (k:l_e_st_eq_landau_n_rt_nt_less yt zt), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (l:l_e_st_eq_landau_n_rt_nt_less xt yt) => (fun (k:l_e_st_eq_landau_n_rt_nt_less yt zt) => l_e_st_eq_landau_n_satz15 (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt) (l_e_st_eq_landau_n_rt_nt_lesse xt yt l) (l_e_st_eq_landau_n_rt_nt_lesse yt zt k)))))). Time Defined. (* constant 1869 *) Definition l_e_st_eq_landau_n_rt_nt_satz15 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), (forall (l:l_e_st_eq_landau_n_rt_nt_less xt yt), (forall (k:l_e_st_eq_landau_n_rt_nt_less yt zt), l_e_st_eq_landau_n_rt_nt_less xt zt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (l:l_e_st_eq_landau_n_rt_nt_less xt yt) => (fun (k:l_e_st_eq_landau_n_rt_nt_less yt zt) => l_e_st_eq_landau_n_rt_nt_lessi xt zt (l_e_st_eq_landau_n_rt_nt_515_t1 xt yt zt l k)))))). Time Defined. (* constant 1870 *) Definition l_e_st_eq_landau_n_rt_nt_521_t1 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ut:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_rt_nt_more xt yt), (forall (n:l_e_st_eq_landau_n_rt_nt_more zt ut), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt ut)))))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ut:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_rt_nt_more xt yt) => (fun (n:l_e_st_eq_landau_n_rt_nt_more zt ut) => l_e_st_eq_landau_n_satz21 (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt zt) (l_e_st_eq_landau_n_rt_nt_nofnt ut) (l_e_st_eq_landau_n_rt_nt_moree xt yt m) (l_e_st_eq_landau_n_rt_nt_moree zt ut n))))))). Time Defined. (* constant 1871 *) Definition l_e_st_eq_landau_n_rt_nt_521_t2 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ut:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_rt_nt_more xt yt), (forall (n:l_e_st_eq_landau_n_rt_nt_more zt ut), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt zt)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt ut)))))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ut:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_rt_nt_more xt yt) => (fun (n:l_e_st_eq_landau_n_rt_nt_more zt ut) => l_e_st_eq_landau_n_ismore12 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt xt) (l_e_st_eq_landau_n_rt_nt_nofnt zt)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl xt zt)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_nt_nofnt yt) (l_e_st_eq_landau_n_rt_nt_nofnt ut)) (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_pl yt ut)) (l_e_st_eq_landau_n_rt_nt_ispln xt zt) (l_e_st_eq_landau_n_rt_nt_ispln yt ut) (l_e_st_eq_landau_n_rt_nt_521_t1 xt yt zt ut m n))))))). Time Defined. (* constant 1872 *) Definition l_e_st_eq_landau_n_rt_nt_satz21 : (forall (xt:l_e_st_eq_landau_n_rt_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_nt_natt), (forall (zt:l_e_st_eq_landau_n_rt_nt_natt), (forall (ut:l_e_st_eq_landau_n_rt_nt_natt), (forall (m:l_e_st_eq_landau_n_rt_nt_more xt yt), (forall (n:l_e_st_eq_landau_n_rt_nt_more zt ut), l_e_st_eq_landau_n_rt_nt_more (l_e_st_eq_landau_n_rt_nt_pl xt zt) (l_e_st_eq_landau_n_rt_nt_pl yt ut))))))). exact (fun (xt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (zt:l_e_st_eq_landau_n_rt_nt_natt) => (fun (ut:l_e_st_eq_landau_n_rt_nt_natt) => (fun (m:l_e_st_eq_landau_n_rt_nt_more xt yt) => (fun (n:l_e_st_eq_landau_n_rt_nt_more zt ut) => l_e_st_eq_landau_n_rt_nt_morei (l_e_st_eq_landau_n_rt_nt_pl xt zt) (l_e_st_eq_landau_n_rt_nt_pl yt ut) (l_e_st_eq_landau_n_rt_nt_521_t2 xt yt zt ut m n))))))). Time Defined. (* constant 1873 *) Definition l_e_st_eq_landau_n_rt_nt_lb : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (n:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (n:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_all (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_imp (p x) (l_e_st_eq_landau_n_rt_nt_lessis n x)))). Time Defined. (* constant 1874 *) Definition l_e_st_eq_landau_n_rt_nt_min : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (n:l_e_st_eq_landau_n_rt_nt_natt), Prop)). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (n:l_e_st_eq_landau_n_rt_nt_natt) => l_and (l_e_st_eq_landau_n_rt_nt_lb p n) (p n))). Time Defined. (* constant 1875 *) Definition l_e_st_eq_landau_n_rt_nt_527_q : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (x:l_e_st_eq_landau_n_nat), Prop))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (x:l_e_st_eq_landau_n_nat) => p (l_e_st_eq_landau_n_rt_nt_ntofn x)))). Time Defined. (* constant 1876 *) Definition l_e_st_eq_landau_n_rt_nt_527_t1 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (n:l_e_st_eq_landau_n_rt_nt_natt), (forall (np:p n), l_e_st_eq_landau_n_rt_nt_527_q p s (l_e_st_eq_landau_n_rt_nt_nofnt n))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (n:l_e_st_eq_landau_n_rt_nt_natt) => (fun (np:p n) => l_e_isp l_e_st_eq_landau_n_rt_nt_natt p n (l_e_st_eq_landau_n_rt_nt_ntofn (l_e_st_eq_landau_n_rt_nt_nofnt n)) np (l_e_st_eq_landau_n_rt_nt_isntn1 n))))). Time Defined. (* constant 1877 *) Definition l_e_st_eq_landau_n_rt_nt_527_t2 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (n:l_e_st_eq_landau_n_rt_nt_natt), (forall (np:p n), l_e_st_eq_landau_n_some (l_e_st_eq_landau_n_rt_nt_527_q p s))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (n:l_e_st_eq_landau_n_rt_nt_natt) => (fun (np:p n) => l_somei l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_nt_527_q p s) (l_e_st_eq_landau_n_rt_nt_nofnt n) (l_e_st_eq_landau_n_rt_nt_527_t1 p s n np))))). Time Defined. (* constant 1878 *) Definition l_e_st_eq_landau_n_rt_nt_527_t3 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), l_e_st_eq_landau_n_some (l_e_st_eq_landau_n_rt_nt_527_q p s))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => l_someapp l_e_st_eq_landau_n_rt_nt_natt p s (l_e_st_eq_landau_n_some (l_e_st_eq_landau_n_rt_nt_527_q p s)) (fun (u:l_e_st_eq_landau_n_rt_nt_natt) => (fun (v:p u) => l_e_st_eq_landau_n_rt_nt_527_t2 p s u v)))). Time Defined. (* constant 1879 *) Definition l_e_st_eq_landau_n_rt_nt_527_t4 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) x))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => l_e_st_eq_landau_n_satz27 (l_e_st_eq_landau_n_rt_nt_527_q p s) (l_e_st_eq_landau_n_rt_nt_527_t3 p s))). Time Defined. (* constant 1880 *) Definition l_e_st_eq_landau_n_rt_nt_527_t5 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), l_e_st_eq_landau_n_lb (l_e_st_eq_landau_n_rt_nt_527_q p s) m)))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => l_ande1 (l_e_st_eq_landau_n_lb (l_e_st_eq_landau_n_rt_nt_527_q p s) m) (l_e_st_eq_landau_n_rt_nt_527_q p s m) mqm)))). Time Defined. (* constant 1881 *) Definition l_e_st_eq_landau_n_rt_nt_527_t6 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), (forall (n:l_e_st_eq_landau_n_nat), (forall (nq:l_e_st_eq_landau_n_rt_nt_527_q p s n), l_e_st_eq_landau_n_lessis m n)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (nq:l_e_st_eq_landau_n_rt_nt_527_q p s n) => l_mp (l_e_st_eq_landau_n_rt_nt_527_q p s n) (l_e_st_eq_landau_n_lessis m n) nq (l_e_st_eq_landau_n_rt_nt_527_t5 p s m mqm n))))))). Time Defined. (* constant 1882 *) Definition l_e_st_eq_landau_n_rt_nt_527_t7 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), (forall (n:l_e_st_eq_landau_n_rt_nt_natt), (forall (np:p n), l_e_st_eq_landau_n_lessis m (l_e_st_eq_landau_n_rt_nt_nofnt n))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => (fun (n:l_e_st_eq_landau_n_rt_nt_natt) => (fun (np:p n) => l_e_st_eq_landau_n_rt_nt_527_t6 p s m mqm (l_e_st_eq_landau_n_rt_nt_nofnt n) (l_e_st_eq_landau_n_rt_nt_527_t1 p s n np))))))). Time Defined. (* constant 1883 *) Definition l_e_st_eq_landau_n_rt_nt_527_t8 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), (forall (n:l_e_st_eq_landau_n_rt_nt_natt), (forall (np:p n), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn m)) (l_e_st_eq_landau_n_rt_nt_nofnt n))))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => (fun (n:l_e_st_eq_landau_n_rt_nt_natt) => (fun (np:p n) => l_e_st_eq_landau_n_islessis1 m (l_e_st_eq_landau_n_rt_nt_nofnt (l_e_st_eq_landau_n_rt_nt_ntofn m)) (l_e_st_eq_landau_n_rt_nt_nofnt n) (l_e_st_eq_landau_n_rt_nt_isnnt1 m) (l_e_st_eq_landau_n_rt_nt_527_t7 p s m mqm n np))))))). Time Defined. (* constant 1884 *) Definition l_e_st_eq_landau_n_rt_nt_527_t9 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), (forall (n:l_e_st_eq_landau_n_rt_nt_natt), (forall (np:p n), l_e_st_eq_landau_n_rt_nt_lessis (l_e_st_eq_landau_n_rt_nt_ntofn m) n)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => (fun (n:l_e_st_eq_landau_n_rt_nt_natt) => (fun (np:p n) => l_e_st_eq_landau_n_rt_nt_lessisi (l_e_st_eq_landau_n_rt_nt_ntofn m) n (l_e_st_eq_landau_n_rt_nt_527_t8 p s m mqm n np))))))). Time Defined. (* constant 1885 *) Definition l_e_st_eq_landau_n_rt_nt_527_t10 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), (forall (n:l_e_st_eq_landau_n_rt_nt_natt), l_imp (p n) (l_e_st_eq_landau_n_rt_nt_lessis (l_e_st_eq_landau_n_rt_nt_ntofn m) n)))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => (fun (n:l_e_st_eq_landau_n_rt_nt_natt) => (fun (u:p n) => l_e_st_eq_landau_n_rt_nt_527_t9 p s m mqm n u)))))). Time Defined. (* constant 1886 *) Definition l_e_st_eq_landau_n_rt_nt_527_t11 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), l_e_st_eq_landau_n_rt_nt_lb p (l_e_st_eq_landau_n_rt_nt_ntofn m))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_527_t10 p s m mqm x))))). Time Defined. (* constant 1887 *) Definition l_e_st_eq_landau_n_rt_nt_527_t12 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), p (l_e_st_eq_landau_n_rt_nt_ntofn m))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => l_ande2 (l_e_st_eq_landau_n_lb (l_e_st_eq_landau_n_rt_nt_527_q p s) m) (l_e_st_eq_landau_n_rt_nt_527_q p s m) mqm)))). Time Defined. (* constant 1888 *) Definition l_e_st_eq_landau_n_rt_nt_527_t13 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), l_e_st_eq_landau_n_rt_nt_min p (l_e_st_eq_landau_n_rt_nt_ntofn m))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => l_andi (l_e_st_eq_landau_n_rt_nt_lb p (l_e_st_eq_landau_n_rt_nt_ntofn m)) (p (l_e_st_eq_landau_n_rt_nt_ntofn m)) (l_e_st_eq_landau_n_rt_nt_527_t11 p s m mqm) (l_e_st_eq_landau_n_rt_nt_527_t12 p s m mqm))))). Time Defined. (* constant 1889 *) Definition l_e_st_eq_landau_n_rt_nt_527_t14 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), (forall (m:l_e_st_eq_landau_n_nat), (forall (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m), l_e_st_eq_landau_n_rt_nt_some (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_min p x))))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (mqm:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) m) => l_somei l_e_st_eq_landau_n_rt_nt_natt (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_min p x) (l_e_st_eq_landau_n_rt_nt_ntofn m) (l_e_st_eq_landau_n_rt_nt_527_t13 p s m mqm))))). Time Defined. (* constant 1890 *) Definition l_e_st_eq_landau_n_rt_nt_satz27 : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)), (forall (s:l_e_st_eq_landau_n_rt_nt_some p), l_e_st_eq_landau_n_rt_nt_some (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_min p x))). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_nt_natt), Prop)) => (fun (s:l_e_st_eq_landau_n_rt_nt_some p) => l_someapp l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) x) (l_e_st_eq_landau_n_rt_nt_527_t4 p s) (l_e_st_eq_landau_n_rt_nt_some (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_min p x)) (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_min (l_e_st_eq_landau_n_rt_nt_527_q p s) x) => l_e_st_eq_landau_n_rt_nt_527_t14 p s x y)))). Time Defined. (* constant 1891 *) Definition l_e_st_eq_landau_n_rt_1rt : l_e_st_eq_landau_n_rt_rat. exact (l_e_st_eq_landau_n_rt_rtofn l_e_st_eq_landau_n_1). Time Defined. (* constant 1892 *) Definition l_e_st_eq_landau_n_rt_ii5_t35 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1)) x))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den x) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) x (l_e_st_eq_landau_n_tfeq1a x l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_eqnd (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x) (l_e_st_eq_landau_n_satz28a (l_e_st_eq_landau_n_num x)) (l_e_st_eq_landau_n_satz28a (l_e_st_eq_landau_n_den x))) (l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) x (l_e_st_eq_landau_n_fris x))))). Time Defined. (* constant 1893 *) Definition l_e_st_eq_landau_n_rt_ii5_t36 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt) x0 (l_e_st_eq_landau_n_tf x (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1)) x (l_e_st_eq_landau_n_rt_tict x0 l_e_st_eq_landau_n_rt_1rt x (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1) xix0 (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1))) xix0 (l_e_st_eq_landau_n_rt_ii5_t35 x0 x xix0)))). Time Defined. (* constant 1894 *) Definition l_e_st_eq_landau_n_rt_example1a : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt) x0). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ratapp1 x0 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt) x0) (fun (x:l_e_st_eq_landau_n_frac) => (fun (xi:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_ii5_t36 x0 x xi))). Time Defined. (* constant 1895 *) Definition l_e_st_eq_landau_n_rt_example1b : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt) x0 (l_e_st_eq_landau_n_rt_example1a x0)). Time Defined. (* constant 1896 *) Definition l_e_st_eq_landau_n_rt_example1c : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x0) x0). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt) x0 (l_e_st_eq_landau_n_rt_comts l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_example1a x0)). Time Defined. (* constant 1897 *) Definition l_e_st_eq_landau_n_rt_example1d : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x0)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x0) x0 (l_e_st_eq_landau_n_rt_example1c x0)). Time Defined. (* constant 1898 *) Definition l_e_st_eq_landau_n_rt_5114_t1 : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_den x) l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_tr3eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_den x) l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den x) (l_e_st_eq_landau_n_num x)) (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_den x))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_den x))) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_tfeq2a x (l_e_st_eq_landau_n_den x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_eqn (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_den x) (l_e_st_eq_landau_n_num x)) (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_den x)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_den x) (l_e_st_eq_landau_n_num x))) (l_e_st_eq_landau_n_satz40c (l_e_st_eq_landau_n_num x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_den x))). Time Defined. (* constant 1899 *) Definition l_e_st_eq_landau_n_rt_satz114 : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_den x)) (l_e_st_eq_landau_n_rt_ratof x)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_num x))). exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_rt_isi (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_den x)) (l_e_st_eq_landau_n_rt_ratof x)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_num x)) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_den x) l_e_st_eq_landau_n_1) x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_tict (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_den x)) (l_e_st_eq_landau_n_rt_ratof x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_den x) l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_den x) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclass x)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_5114_t1 x)). Time Defined. (* constant 1900 *) Definition l_e_st_eq_landau_n_rt_satz114a : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x2) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_rt_rtofn x1))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x2) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_rt_rtofn x1) (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_rtofn x2) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2))) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_rt_isnert x2 (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_isden x1 x2))) (l_e_st_eq_landau_n_rt_satz114 (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_rt_isnert (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr x1 x2)) x1 (l_e_st_eq_landau_n_numis x1 x2)))). Time Defined. (* constant 1901 *) Definition l_e_st_eq_landau_n_rt_ov : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rat)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_ind l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0) (l_e_st_eq_landau_n_rt_satz110 x0 y0))). Time Defined. (* constant 1902 *) Definition l_e_st_eq_landau_n_rt_satz110c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov x0 y0)) x0)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_oneax l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 t) x0) (l_e_st_eq_landau_n_rt_satz110 x0 y0))). Time Defined. (* constant 1903 *) Definition l_e_st_eq_landau_n_rt_satz110d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov x0 y0)) x0 (l_e_st_eq_landau_n_rt_satz110c x0 y0))). Time Defined. (* constant 1904 *) Definition l_e_st_eq_landau_n_rt_satz110e : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov x0 y0) y0) x0)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov x0 y0) y0) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov x0 y0)) x0 (l_e_st_eq_landau_n_rt_comts (l_e_st_eq_landau_n_rt_ov x0 y0) y0) (l_e_st_eq_landau_n_rt_satz110c x0 y0))). Time Defined. (* constant 1905 *) Definition l_e_st_eq_landau_n_rt_satz110f : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov x0 y0) y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov x0 y0) y0) x0 (l_e_st_eq_landau_n_rt_satz110e x0 y0))). Time Defined. (* constant 1906 *) Definition l_e_st_eq_landau_n_rt_satz110g : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 v0) x0), l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ov x0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 v0) x0) => l_e_st_eq_landau_n_rt_satz110b x0 y0 v0 (l_e_st_eq_landau_n_rt_ov x0 y0) i (l_e_st_eq_landau_n_rt_satz110c x0 y0))))). Time Defined. (* constant 1907 *) Definition l_e_st_eq_landau_n_rt_satz114b : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn x1) (l_e_st_eq_landau_n_rt_rtofn x2)))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_satz110b (l_e_st_eq_landau_n_rt_rtofn x1) (l_e_st_eq_landau_n_rt_rtofn x2) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn x1) (l_e_st_eq_landau_n_rt_rtofn x2)) (l_e_st_eq_landau_n_rt_satz114a x1 x2) (l_e_st_eq_landau_n_rt_satz110c (l_e_st_eq_landau_n_rt_rtofn x1) (l_e_st_eq_landau_n_rt_rtofn x2)))). Time Defined. (* constant 1908 *) Definition l_e_st_eq_landau_n_rt_satz114c : (forall (x1:l_e_st_eq_landau_n_nat), (forall (x2:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn x1) (l_e_st_eq_landau_n_rt_rtofn x2)) (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x1 x2)))). exact (fun (x1:l_e_st_eq_landau_n_nat) => (fun (x2:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr x1 x2)) (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn x1) (l_e_st_eq_landau_n_rt_rtofn x2)) (l_e_st_eq_landau_n_rt_satz114b x1 x2))). Time Defined. (* constant 1909 *) Definition l_e_st_eq_landau_n_rt_5115_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t (l_e_st_eq_landau_n_rt_ov y0 x0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz89 (l_e_st_eq_landau_n_rt_ov y0 x0))). Time Defined. (* constant 1910 *) Definition l_e_st_eq_landau_n_rt_5115_z : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_nat)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_num u)))))). Time Defined. (* constant 1911 *) Definition l_e_st_eq_landau_n_rt_5115_v : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_nat)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_den u)))))). Time Defined. (* constant 1912 *) Definition l_e_st_eq_landau_n_rt_5115_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_ov y0 x0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_ismore1 u0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_ov y0 x0) (l_e_tris l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_isi u0 (l_e_st_eq_landau_n_rt_ratof (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) u (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) uiu0 (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_refeq1 u (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_isfr u))) (l_e_st_eq_landau_n_rt_satz114b (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) m)))))). Time Defined. (* constant 1913 *) Definition l_e_st_eq_landau_n_rt_5115_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_moreisi (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1)) (l_or_th9 (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_satz24 (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) (fun (t:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_satz111d (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1 t) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_satz111e (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0) l_e_st_eq_landau_n_1 t)))))))). Time Defined. (* constant 1914 *) Definition l_e_st_eq_landau_n_rt_5115_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_comts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ists2 (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) x0 (l_e_st_eq_landau_n_rt_satz110f (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) (l_e_st_eq_landau_n_rt_assts2 x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))))))))). Time Defined. (* constant 1915 *) Definition l_e_st_eq_landau_n_rt_5115_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), (forall (n:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => (fun (n:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt) => l_e_st_eq_landau_n_rt_ismore1 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt) (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_5115_t4 x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_satz105d (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) n)))))))). Time Defined. (* constant 1916 *) Definition l_e_st_eq_landau_n_rt_5115_t6 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), (forall (n:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt), l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => (fun (n:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt) => l_e_st_eq_landau_n_rt_moreisi1 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt) (l_e_st_eq_landau_n_rt_5115_t5 x0 y0 u0 m u uiu0 n)))))))). Time Defined. (* constant 1917 *) Definition l_e_st_eq_landau_n_rt_5115_t7 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt), l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt) => l_e_st_eq_landau_n_rt_moreisi2 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt) (l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt) (l_e_st_eq_landau_n_rt_5115_t4 x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_ists2 (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) i))))))))). Time Defined. (* constant 1918 *) Definition l_e_st_eq_landau_n_rt_5115_t8 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_orapp (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt) (l_e_st_eq_landau_n_rt_moreis (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt)) (l_e_st_eq_landau_n_rt_5115_t3 x0 y0 u0 m u uiu0) (fun (t:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt) => l_e_st_eq_landau_n_rt_5115_t6 x0 y0 u0 m u uiu0 t) (fun (t:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)) l_e_st_eq_landau_n_rt_1rt) => l_e_st_eq_landau_n_rt_5115_t7 x0 y0 u0 m u uiu0 t))))))). Time Defined. (* constant 1919 *) Definition l_e_st_eq_landau_n_rt_5115_t9 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt) y0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_ismore12 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov y0 x0)) y0 (l_e_st_eq_landau_n_rt_example1b (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))))) (l_e_st_eq_landau_n_rt_satz110c y0 x0) (l_e_st_eq_landau_n_rt_satz105d (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_ov y0 x0) x0 (l_e_st_eq_landau_n_rt_5115_t2 x0 y0 u0 m u uiu0)))))))). Time Defined. (* constant 1920 *) Definition l_e_st_eq_landau_n_rt_5115_t10 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) y0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_satz87c (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_v x0 y0 u0 m u uiu0)))) l_e_st_eq_landau_n_rt_1rt) y0 (l_e_st_eq_landau_n_rt_5115_t8 x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_5115_t9 x0 y0 u0 m u uiu0))))))). Time Defined. (* constant 1921 *) Definition l_e_st_eq_landau_n_rt_5115_t11 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_some (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn t) x0) y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn t) x0) y0) (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0) (l_e_st_eq_landau_n_rt_5115_t10 x0 y0 u0 m u uiu0))))))). Time Defined. (* constant 1922 *) Definition l_e_st_eq_landau_n_rt_5115_t12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), l_e_st_eq_landau_n_some (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn t) x0) y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => l_e_st_eq_landau_n_rt_ratapp1 u0 (l_e_st_eq_landau_n_some (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn t) x0) y0)) (fun (u:l_e_st_eq_landau_n_frac) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_5115_t11 x0 y0 u0 m u ui)))))). Time Defined. (* constant 1923 *) Definition l_e_st_eq_landau_n_rt_satz115 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_some (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn t) x0) y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t (l_e_st_eq_landau_n_rt_ov y0 x0)) (l_e_st_eq_landau_n_rt_5115_t1 x0 y0) (l_e_st_eq_landau_n_some (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn t) x0) y0)) (fun (t:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_more t (l_e_st_eq_landau_n_rt_ov y0 x0)) => l_e_st_eq_landau_n_rt_5115_t12 x0 y0 t u)))). Time Defined. (* constant 1924 *) Definition l_e_st_eq_landau_n_rt_5115_t13 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_and (l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_andi (l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0))) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) x0) y0) (l_e_st_eq_landau_n_rt_natrti (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_5115_t10 x0 y0 u0 m u uiu0))))))). Time Defined. (* constant 1925 *) Definition l_e_st_eq_landau_n_rt_5115_t14 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), (forall (u:l_e_st_eq_landau_n_frac), (forall (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_natrt t) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts t x0) y0)))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => (fun (u:l_e_st_eq_landau_n_frac) => (fun (uiu0:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_natrt t) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts t x0) y0)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_5115_z x0 y0 u0 m u uiu0)) (l_e_st_eq_landau_n_rt_5115_t13 x0 y0 u0 m u uiu0))))))). Time Defined. (* constant 1926 *) Definition l_e_st_eq_landau_n_rt_5115_t15 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_natrt t) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts t x0) y0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more u0 (l_e_st_eq_landau_n_rt_ov y0 x0)) => l_e_st_eq_landau_n_rt_ratapp1 u0 (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_natrt t) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts t x0) y0))) (fun (u:l_e_st_eq_landau_n_frac) => (fun (ui:l_e_st_eq_landau_n_rt_inf u (l_e_st_eq_landau_n_rt_class u0)) => l_e_st_eq_landau_n_rt_5115_t14 x0 y0 u0 m u ui)))))). Time Defined. (* constant 1927 *) Definition l_e_st_eq_landau_n_rt_satz115a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_natrt t) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts t x0) y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_more t (l_e_st_eq_landau_n_rt_ov y0 x0)) (l_e_st_eq_landau_n_rt_5115_t1 x0 y0) (l_e_st_eq_landau_n_rt_some (fun (t:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_natrt t) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts t x0) y0))) (fun (t:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_more t (l_e_st_eq_landau_n_rt_ov y0 x0)) => l_e_st_eq_landau_n_rt_5115_t15 x0 y0 t u)))). Time Defined. (* constant 1928 *) Definition l_e_st_eq_landau_n_rt_cutprop1a : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_nonempty l_e_st_eq_landau_n_rt_rat s). Time Defined. (* constant 1929 *) Definition l_e_st_eq_landau_n_rt_cutprop1b : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_not (l_e_st_eq_landau_n_rt_all (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_in x s))). Time Defined. (* constant 1930 *) Definition l_e_st_eq_landau_n_rt_cutprop1 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_cutprop1a s) (l_e_st_eq_landau_n_rt_cutprop1b s)). Time Defined. (* constant 1931 *) Definition l_e_st_eq_landau_n_rt_cutprop2a : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_all (fun (x:l_e_st_eq_landau_n_rt_rat) => l_imp (l_not (l_e_st_eq_landau_n_rt_in x s)) (l_e_st_eq_landau_n_rt_less x0 x)))). Time Defined. (* constant 1932 *) Definition l_e_st_eq_landau_n_rt_cutprop2 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_all (fun (x:l_e_st_eq_landau_n_rt_rat) => l_imp (l_e_st_eq_landau_n_rt_in x s) (l_e_st_eq_landau_n_rt_cutprop2a s x))). Time Defined. (* constant 1933 *) Definition l_e_st_eq_landau_n_rt_iii1_ubprop : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_imp (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_moreis x0 y0)))). Time Defined. (* constant 1934 *) Definition l_e_st_eq_landau_n_rt_ub : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_all (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_iii1_ubprop s x0 x))). Time Defined. (* constant 1935 *) Definition l_e_st_eq_landau_n_rt_max : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_ub s x0) (l_e_st_eq_landau_n_rt_in x0 s))). Time Defined. (* constant 1936 *) Definition l_e_st_eq_landau_n_rt_cutprop3 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_not (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_max s x))). Time Defined. (* constant 1937 *) Definition l_e_st_eq_landau_n_rt_cutprop : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_and3 (l_e_st_eq_landau_n_rt_cutprop1 s) (l_e_st_eq_landau_n_rt_cutprop2 s) (l_e_st_eq_landau_n_rt_cutprop3 s)). Time Defined. (* constant 1938 *) Definition l_e_st_eq_landau_n_rt_iii1_lbprop : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_imp (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_lessis x0 y0)))). Time Defined. (* constant 1939 *) Definition l_e_st_eq_landau_n_rt_lb : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_all (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_iii1_lbprop s x0 x))). Time Defined. (* constant 1940 *) Definition l_e_st_eq_landau_n_rt_min : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lb s x0) (l_e_st_eq_landau_n_rt_in x0 s))). Time Defined. (* constant 1941 *) Definition l_e_st_eq_landau_n_rt_cut : Type. exact (l_e_ot (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x)). Time Defined. (* constant 1942 *) Definition l_e_st_eq_landau_n_rt_lcl : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_set l_e_st_eq_landau_n_rt_rat). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_in (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x) ksi). Time Defined. (* constant 1943 *) Definition l_e_st_eq_landau_n_rt_lrt : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_lcl ksi))). Time Defined. (* constant 1944 *) Definition l_e_st_eq_landau_n_rt_urt : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_not (l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_lcl ksi)))). Time Defined. (* constant 1945 *) Definition l_e_st_eq_landau_n_rt_clcl : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_lcl ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_inp (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x) ksi). Time Defined. (* constant 1946 *) Definition l_e_st_eq_landau_n_rt_clcl1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop1 (l_e_st_eq_landau_n_rt_lcl ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_and3e1 (l_e_st_eq_landau_n_rt_cutprop1 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_cutprop2 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_cutprop3 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_clcl ksi)). Time Defined. (* constant 1947 *) Definition l_e_st_eq_landau_n_rt_clcl2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop2 (l_e_st_eq_landau_n_rt_lcl ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_and3e2 (l_e_st_eq_landau_n_rt_cutprop1 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_cutprop2 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_cutprop3 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_clcl ksi)). Time Defined. (* constant 1948 *) Definition l_e_st_eq_landau_n_rt_clcl3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop3 (l_e_st_eq_landau_n_rt_lcl ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_and3e3 (l_e_st_eq_landau_n_rt_cutprop1 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_cutprop2 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_cutprop3 (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_clcl ksi)). Time Defined. (* constant 1949 *) Definition l_e_st_eq_landau_n_rt_clcl1a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop1a (l_e_st_eq_landau_n_rt_lcl ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_ande1 (l_e_st_eq_landau_n_rt_cutprop1a (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_cutprop1b (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_clcl1 ksi)). Time Defined. (* constant 1950 *) Definition l_e_st_eq_landau_n_rt_clcl1b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop1b (l_e_st_eq_landau_n_rt_lcl ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_ande2 (l_e_st_eq_landau_n_rt_cutprop1a (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_cutprop1b (l_e_st_eq_landau_n_rt_lcl ksi)) (l_e_st_eq_landau_n_rt_clcl1 ksi)). Time Defined. (* constant 1951 *) Definition l_e_st_eq_landau_n_rt_cutapp1a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), p))), p))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), p))) => l_e_st_nonemptyapp l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_clcl1a ksi) p p1))). Time Defined. (* constant 1952 *) Definition l_e_st_eq_landau_n_rt_iii1_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_some_th1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lrt ksi x) (l_e_st_eq_landau_n_rt_clcl1b ksi)). Time Defined. (* constant 1953 *) Definition l_e_st_eq_landau_n_rt_cutapp1b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), p))), p))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), p))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x) (l_e_st_eq_landau_n_rt_iii1_t1 ksi) p p1))). Time Defined. (* constant 1954 *) Definition l_e_st_eq_landau_n_rt_iii1_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_cutprop2a (l_e_st_eq_landau_n_rt_lcl ksi) x0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_mp (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_cutprop2a (l_e_st_eq_landau_n_rt_lcl ksi) x0) lx (l_e_st_eq_landau_n_rt_clcl2 ksi x0)))). Time Defined. (* constant 1955 *) Definition l_e_st_eq_landau_n_rt_cutapp2a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_rt_less x0 y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_mp (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_less x0 y0) uy (l_e_st_eq_landau_n_rt_iii1_t2 ksi x0 lx y0)))))). Time Defined. (* constant 1956 *) Definition l_e_st_eq_landau_n_rt_cutapp2b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_rt_more y0 x0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_e_st_eq_landau_n_rt_satz83 x0 y0 (l_e_st_eq_landau_n_rt_cutapp2a ksi x0 lx y0 uy)))))). Time Defined. (* constant 1957 *) Definition l_e_st_eq_landau_n_rt_iii1_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_not (l_e_st_eq_landau_n_rt_max (l_e_st_eq_landau_n_rt_lcl ksi) x0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_some_th4 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_max (l_e_st_eq_landau_n_rt_lcl ksi) x) (l_e_st_eq_landau_n_rt_clcl3 ksi) x0))). Time Defined. (* constant 1958 *) Definition l_e_st_eq_landau_n_rt_iii1_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_not (l_e_st_eq_landau_n_rt_ub (l_e_st_eq_landau_n_rt_lcl ksi) x0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_and_th4 (l_e_st_eq_landau_n_rt_ub (l_e_st_eq_landau_n_rt_lcl ksi) x0) (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_iii1_t3 ksi x0 lx) lx))). Time Defined. (* constant 1959 *) Definition l_e_st_eq_landau_n_rt_iii1_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 x))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_some_th1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 x) (l_e_st_eq_landau_n_rt_iii1_t4 ksi x0 lx)))). Time Defined. (* constant 1960 *) Definition l_e_st_eq_landau_n_rt_iii1_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (p:Prop), (forall (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y0)), l_e_st_eq_landau_n_rt_lrt ksi y0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (p:Prop) => (fun (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y0)) => l_imp_th5 (l_e_st_eq_landau_n_rt_lrt ksi y0) (l_e_st_eq_landau_n_rt_moreis x0 y0) n))))))). Time Defined. (* constant 1961 *) Definition l_e_st_eq_landau_n_rt_iii1_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (p:Prop), (forall (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y0)), l_not (l_e_st_eq_landau_n_rt_moreis x0 y0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (p:Prop) => (fun (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y0)) => l_imp_th6 (l_e_st_eq_landau_n_rt_lrt ksi y0) (l_e_st_eq_landau_n_rt_moreis x0 y0) n))))))). Time Defined. (* constant 1962 *) Definition l_e_st_eq_landau_n_rt_iii1_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (p:Prop), (forall (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y0)), l_e_st_eq_landau_n_rt_less x0 y0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (p:Prop) => (fun (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y0)) => l_e_st_eq_landau_n_rt_satz81j x0 y0 (l_e_st_eq_landau_n_rt_iii1_t7 ksi x0 lx p p1 y0 n)))))))). Time Defined. (* constant 1963 *) Definition l_e_st_eq_landau_n_rt_iii1_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (p:Prop), (forall (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y0)), p))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (p:Prop) => (fun (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y0)) => p1 y0 (l_e_st_eq_landau_n_rt_iii1_t6 ksi x0 lx p p1 y0 n) (l_e_st_eq_landau_n_rt_iii1_t8 ksi x0 lx p p1 y0 n)))))))). Time Defined. (* constant 1964 *) Definition l_e_st_eq_landau_n_rt_cutapp3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (p:Prop), (forall (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))), p))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (p:Prop) => (fun (p1:(forall (y:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi y), (forall (u:l_e_st_eq_landau_n_rt_less x0 y), p)))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y)) (l_e_st_eq_landau_n_rt_iii1_t5 ksi x0 lx) p (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_not (l_e_st_eq_landau_n_rt_iii1_ubprop (l_e_st_eq_landau_n_rt_lcl ksi) x0 y)) => l_e_st_eq_landau_n_rt_iii1_t9 ksi x0 lx p p1 y t))))))). Time Defined. (* constant 1965 *) Definition l_e_st_eq_landau_n_rt_iii1_t10 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), l_e_st_eq_landau_n_rt_cutprop1 s))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => l_andi (l_e_st_eq_landau_n_rt_cutprop1a s) (l_e_st_eq_landau_n_rt_cutprop1b s) (l_e_st_nonemptyi l_e_st_eq_landau_n_rt_rat s x0 i) (l_all_th1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_in x s) y0 n)))))). Time Defined. (* constant 1966 *) Definition l_e_st_eq_landau_n_rt_iii1_t11 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (n:(forall (x:l_e_st_eq_landau_n_rt_rat), l_not (l_e_st_eq_landau_n_rt_max s x))), l_e_st_eq_landau_n_rt_cutprop3 s)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (n:(forall (x:l_e_st_eq_landau_n_rt_rat), l_not (l_e_st_eq_landau_n_rt_max s x))) => l_some_th5 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_max s x) n)). Time Defined. (* constant 1967 *) Definition l_e_st_eq_landau_n_rt_cut1 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), (forall (l:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_not (l_e_st_eq_landau_n_rt_in y s)), l_e_st_eq_landau_n_rt_less x y))))), (forall (m:(forall (x:l_e_st_eq_landau_n_rt_rat), l_not (l_e_st_eq_landau_n_rt_max s x))), l_e_st_eq_landau_n_rt_cutprop s))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => (fun (l:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_not (l_e_st_eq_landau_n_rt_in y s)), l_e_st_eq_landau_n_rt_less x y))))) => (fun (m:(forall (x:l_e_st_eq_landau_n_rt_rat), l_not (l_e_st_eq_landau_n_rt_max s x))) => l_and3i (l_e_st_eq_landau_n_rt_cutprop1 s) (l_e_st_eq_landau_n_rt_cutprop2 s) (l_e_st_eq_landau_n_rt_cutprop3 s) (l_e_st_eq_landau_n_rt_iii1_t10 s x0 i y0 n) l (l_e_st_eq_landau_n_rt_iii1_t11 s m)))))))). Time Defined. (* constant 1968 *) Definition l_e_st_eq_landau_n_rt_rp_is : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_is l_e_st_eq_landau_n_rt_cut ksi eta)). Time Defined. (* constant 1969 *) Definition l_e_st_eq_landau_n_rt_rp_nis : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_not (l_e_st_eq_landau_n_rt_rp_is ksi eta))). Time Defined. (* constant 1970 *) Definition l_e_st_eq_landau_n_rt_rp_ise : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isini (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x) ksi eta i))). Time Defined. (* constant 1971 *) Definition l_e_st_eq_landau_n_rt_rp_ise1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_lrt eta x0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_issete1 l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta) (l_e_st_eq_landau_n_rt_rp_ise ksi eta i) x0 lx))))). Time Defined. (* constant 1972 *) Definition l_e_st_eq_landau_n_rt_rp_isi : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)), l_e_st_eq_landau_n_rt_rp_is ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)) => l_e_isine (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x) ksi eta i))). Time Defined. (* constant 1973 *) Definition l_e_st_eq_landau_n_rt_rp_isi1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), l_e_st_eq_landau_n_rt_lrt eta x))), (forall (k:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt eta x), l_e_st_eq_landau_n_rt_lrt ksi x))), l_e_st_eq_landau_n_rt_rp_is ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), l_e_st_eq_landau_n_rt_lrt eta x))) => (fun (k:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt eta x), l_e_st_eq_landau_n_rt_lrt ksi x))) => l_e_st_eq_landau_n_rt_rp_isi ksi eta (l_e_st_isseti l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta) l k))))). Time Defined. (* constant 1974 *) Definition l_e_st_eq_landau_n_rt_rp_cutof : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (cs:l_e_st_eq_landau_n_rt_cutprop s), l_e_st_eq_landau_n_rt_cut)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (cs:l_e_st_eq_landau_n_rt_cutprop s) => l_e_out (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x) s cs)). Time Defined. (* constant 1975 *) Definition l_e_st_eq_landau_n_rt_rp_ine : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (cs:l_e_st_eq_landau_n_rt_cutprop s), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_cutof s cs) x0)))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (cs:l_e_st_eq_landau_n_rt_cutprop s) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => l_e_st_issete1 l_e_st_eq_landau_n_rt_rat s (l_e_st_eq_landau_n_rt_lcl (l_e_st_eq_landau_n_rt_rp_cutof s cs)) (l_e_isinout (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x) s cs) x0 i)))). Time Defined. (* constant 1976 *) Definition l_e_st_eq_landau_n_rt_rp_ini : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (cs:l_e_st_eq_landau_n_rt_cutprop s), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_cutof s cs) x0), l_e_st_eq_landau_n_rt_in x0 s)))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (cs:l_e_st_eq_landau_n_rt_cutprop s) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_cutof s cs) x0) => l_e_st_issete2 l_e_st_eq_landau_n_rt_rat s (l_e_st_eq_landau_n_rt_lcl (l_e_st_eq_landau_n_rt_rp_cutof s cs)) (l_e_isinout (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x) s cs) x0 lx)))). Time Defined. (* constant 1977 *) Definition l_e_st_eq_landau_n_rt_rp_isi2 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (cs:l_e_st_eq_landau_n_rt_cutprop s), (forall (ct:l_e_st_eq_landau_n_rt_cutprop t), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_in x t))), (forall (j:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_in x t), l_e_st_eq_landau_n_rt_in x s))), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_cutof s cs) (l_e_st_eq_landau_n_rt_rp_cutof t ct))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (cs:l_e_st_eq_landau_n_rt_cutprop s) => (fun (ct:l_e_st_eq_landau_n_rt_cutprop t) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_in x t))) => (fun (j:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_in x t), l_e_st_eq_landau_n_rt_in x s))) => l_e_isouti (l_e_st_set l_e_st_eq_landau_n_rt_rat) (fun (x:l_e_st_set l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutprop x) s cs t ct (l_e_st_isseti l_e_st_eq_landau_n_rt_rat s t i j))))))). Time Defined. (* constant 1978 *) Definition l_e_st_eq_landau_n_rt_rp_all : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)) => l_all l_e_st_eq_landau_n_rt_cut p). Time Defined. (* constant 1979 *) Definition l_e_st_eq_landau_n_rt_rp_some : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)) => l_some l_e_st_eq_landau_n_rt_cut p). Time Defined. (* constant 1980 *) Definition l_e_st_eq_landau_n_rt_rp_one : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)) => l_e_one l_e_st_eq_landau_n_rt_cut p). Time Defined. (* constant 1981 *) Definition l_e_st_eq_landau_n_rt_rp_satz116 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is ksi ksi). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_refis l_e_st_eq_landau_n_rt_cut ksi). Time Defined. (* constant 1982 *) Definition l_e_st_eq_landau_n_rt_rp_satz117 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_symis l_e_st_eq_landau_n_rt_cut ksi eta i))). Time Defined. (* constant 1983 *) Definition l_e_st_eq_landau_n_rt_rp_satz118 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is eta zeta), l_e_st_eq_landau_n_rt_rp_is ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is eta zeta) => l_e_tris l_e_st_eq_landau_n_rt_cut ksi eta zeta i j))))). Time Defined. (* constant 1984 *) Definition l_e_st_eq_landau_n_rt_rp_1119_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_not (l_e_st_eq_landau_n_rt_less x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_ec3e23 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) m))). Time Defined. (* constant 1985 *) Definition l_e_st_eq_landau_n_rt_rp_satz119 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more y0 x0), l_e_st_eq_landau_n_rt_urt ksi y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more y0 x0) => l_imp_th3 (l_e_st_eq_landau_n_rt_lrt ksi y0) (l_e_st_eq_landau_n_rt_less y0 x0) (l_e_st_eq_landau_n_rt_rp_1119_t1 y0 x0 m) (fun (t:l_e_st_eq_landau_n_rt_lrt ksi y0) => l_e_st_eq_landau_n_rt_cutapp2a ksi y0 t x0 ux)))))). Time Defined. (* constant 1986 *) Definition l_e_st_eq_landau_n_rt_rp_satz119a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_urt ksi y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_rp_satz119 ksi x0 ux y0 (l_e_st_eq_landau_n_rt_satz83 x0 y0 l)))))). Time Defined. (* constant 1987 *) Definition l_e_st_eq_landau_n_rt_rp_1120_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_not (l_e_st_eq_landau_n_rt_more x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_ec3e32 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) l))). Time Defined. (* constant 1988 *) Definition l_e_st_eq_landau_n_rt_rp_satz120 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_lrt ksi y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_imp_th7 (l_e_st_eq_landau_n_rt_lrt ksi y0) (l_e_st_eq_landau_n_rt_more y0 x0) (l_e_st_eq_landau_n_rt_rp_1120_t1 y0 x0 l) (fun (t:l_e_st_eq_landau_n_rt_urt ksi y0) => l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx y0 t)))))). Time Defined. (* constant 1989 *) Definition l_e_st_eq_landau_n_rt_rp_satz120a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_lrt ksi y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_rp_satz120 ksi x0 lx y0 (l_e_st_eq_landau_n_rt_satz82 x0 y0 m)))))). Time Defined. (* constant 1990 *) Definition l_e_st_eq_landau_n_rt_iii1_t12 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), (forall (u:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_in y0 s))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => i x0 j y0)))))). Time Defined. (* constant 1991 *) Definition l_e_st_eq_landau_n_rt_iii1_t13 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), l_not (l_e_st_eq_landau_n_rt_less y0 x0))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => l_imp_th3 (l_e_st_eq_landau_n_rt_less y0 x0) (l_e_st_eq_landau_n_rt_in y0 s) n (l_e_st_eq_landau_n_rt_iii1_t12 s i x0 j y0 n))))))). Time Defined. (* constant 1992 *) Definition l_e_st_eq_landau_n_rt_iii1_t14 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), l_e_st_eq_landau_n_rt_moreis y0 x0)))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => l_e_st_eq_landau_n_rt_satz81f y0 x0 (l_e_st_eq_landau_n_rt_iii1_t13 s i x0 j y0 n))))))). Time Defined. (* constant 1993 *) Definition l_e_st_eq_landau_n_rt_iii1_t15 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), (forall (k:l_e_st_eq_landau_n_rt_is y0 x0), l_e_st_eq_landau_n_rt_in y0 s))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => (fun (k:l_e_st_eq_landau_n_rt_is y0 x0) => l_e_isp1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_in x s) x0 y0 j k))))))). Time Defined. (* constant 1994 *) Definition l_e_st_eq_landau_n_rt_iii1_t16 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), l_e_st_eq_landau_n_rt_nis y0 x0)))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => l_imp_th3 (l_e_st_eq_landau_n_rt_is y0 x0) (l_e_st_eq_landau_n_rt_in y0 s) n (fun (t:l_e_st_eq_landau_n_rt_is y0 x0) => l_e_st_eq_landau_n_rt_iii1_t15 s i x0 j y0 n t))))))). Time Defined. (* constant 1995 *) Definition l_e_st_eq_landau_n_rt_iii1_t17 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), l_e_st_eq_landau_n_rt_more y0 x0)))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => l_ore1 (l_e_st_eq_landau_n_rt_more y0 x0) (l_e_st_eq_landau_n_rt_is y0 x0) (l_e_st_eq_landau_n_rt_iii1_t14 s i x0 j y0 n) (l_e_st_eq_landau_n_rt_iii1_t16 s i x0 j y0 n))))))). Time Defined. (* constant 1996 *) Definition l_e_st_eq_landau_n_rt_iii1_t18 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), l_e_st_eq_landau_n_rt_less x0 y0)))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => l_e_st_eq_landau_n_rt_satz82 y0 x0 (l_e_st_eq_landau_n_rt_iii1_t17 s i x0 j y0 n))))))). Time Defined. (* constant 1997 *) Definition l_e_st_eq_landau_n_rt_iii1_t19 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), l_e_st_eq_landau_n_rt_cutprop2 s)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (i:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x s) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_not (l_e_st_eq_landau_n_rt_in y s)) => l_e_st_eq_landau_n_rt_iii1_t18 s i x t y u)))))). Time Defined. (* constant 1998 *) Definition l_e_st_eq_landau_n_rt_iii1_t20 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x0)))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => s1 x0 i)))). Time Defined. (* constant 1999 *) Definition l_e_st_eq_landau_n_rt_iii1_t21 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)), l_e_st_eq_landau_n_rt_in y0 s)))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)) => l_ande1 (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0) a)))))). Time Defined. (* constant 2000 *) Definition l_e_st_eq_landau_n_rt_iii1_t22 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)), l_e_st_eq_landau_n_rt_more y0 x0)))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)) => l_ande2 (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0) a)))))). Time Defined. (* constant 2001 *) Definition l_e_st_eq_landau_n_rt_iii1_t23 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)), l_not (l_e_st_eq_landau_n_rt_lessis y0 x0))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)) => l_e_st_eq_landau_n_rt_satz81g y0 x0 (l_e_st_eq_landau_n_rt_iii1_t22 s s1 x0 i y0 a))))))). Time Defined. (* constant 2002 *) Definition l_e_st_eq_landau_n_rt_iii1_t24 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)), l_not (l_e_st_eq_landau_n_rt_moreis x0 y0))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)) => l_imp_th3 (l_e_st_eq_landau_n_rt_moreis x0 y0) (l_e_st_eq_landau_n_rt_lessis y0 x0) (l_e_st_eq_landau_n_rt_iii1_t23 s s1 x0 i y0 a) (fun (t:l_e_st_eq_landau_n_rt_moreis x0 y0) => l_e_st_eq_landau_n_rt_satz84 x0 y0 t))))))). Time Defined. (* constant 2003 *) Definition l_e_st_eq_landau_n_rt_iii1_t25 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)), l_not (l_e_st_eq_landau_n_rt_iii1_ubprop s x0 y0))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)) => l_imp_th4 (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_moreis x0 y0) (l_e_st_eq_landau_n_rt_iii1_t21 s s1 x0 i y0 a) (l_e_st_eq_landau_n_rt_iii1_t24 s s1 x0 i y0 a))))))). Time Defined. (* constant 2004 *) Definition l_e_st_eq_landau_n_rt_iii1_t26 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)), l_not (l_e_st_eq_landau_n_rt_ub s x0))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)) => l_all_th1 l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_iii1_ubprop s x0 y) y0 (l_e_st_eq_landau_n_rt_iii1_t25 s s1 x0 i y0 a))))))). Time Defined. (* constant 2005 *) Definition l_e_st_eq_landau_n_rt_iii1_t27 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)), l_not (l_e_st_eq_landau_n_rt_max s x0))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_in y0 s) (l_e_st_eq_landau_n_rt_more y0 x0)) => l_and_th1 (l_e_st_eq_landau_n_rt_ub s x0) (l_e_st_eq_landau_n_rt_in x0 s) (l_e_st_eq_landau_n_rt_iii1_t26 s s1 x0 i y0 a))))))). Time Defined. (* constant 2006 *) Definition l_e_st_eq_landau_n_rt_iii1_t28 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), l_not (l_e_st_eq_landau_n_rt_max s x0))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x0)) (l_e_st_eq_landau_n_rt_iii1_t20 s s1 x0 i) (l_not (l_e_st_eq_landau_n_rt_max s x0)) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x0)) => l_e_st_eq_landau_n_rt_iii1_t27 s s1 x0 i y t)))))). Time Defined. (* constant 2007 *) Definition l_e_st_eq_landau_n_rt_iii1_t29 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in x0 s)), l_not (l_e_st_eq_landau_n_rt_max s x0))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in x0 s)) => l_and_th2 (l_e_st_eq_landau_n_rt_ub s x0) (l_e_st_eq_landau_n_rt_in x0 s) n)))). Time Defined. (* constant 2008 *) Definition l_e_st_eq_landau_n_rt_iii1_t30 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), l_not (l_e_st_eq_landau_n_rt_max s x0)))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_imp_th1 (l_e_st_eq_landau_n_rt_in x0 s) (l_not (l_e_st_eq_landau_n_rt_max s x0)) (fun (u:l_e_st_eq_landau_n_rt_in x0 s) => l_e_st_eq_landau_n_rt_iii1_t28 s s1 x0 u) (fun (u:l_not (l_e_st_eq_landau_n_rt_in x0 s)) => l_e_st_eq_landau_n_rt_iii1_t29 s s1 x0 u)))). Time Defined. (* constant 2009 *) Definition l_e_st_eq_landau_n_rt_iii1_t31 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), l_e_st_eq_landau_n_rt_cutprop3 s)). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => l_e_st_eq_landau_n_rt_iii1_t11 s (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_iii1_t30 s s1 x))). Time Defined. (* constant 2010 *) Definition l_e_st_eq_landau_n_rt_cut2 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 s), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)), (forall (j:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))), (forall (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))), l_e_st_eq_landau_n_rt_cutprop s))))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 s) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (n:l_not (l_e_st_eq_landau_n_rt_in y0 s)) => (fun (j:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_less y x), l_e_st_eq_landau_n_rt_in y s))))) => (fun (s1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_in x s), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y s) (l_e_st_eq_landau_n_rt_more y x))))) => l_and3i (l_e_st_eq_landau_n_rt_cutprop1 s) (l_e_st_eq_landau_n_rt_cutprop2 s) (l_e_st_eq_landau_n_rt_cutprop3 s) (l_e_st_eq_landau_n_rt_iii1_t10 s x0 i y0 n) (l_e_st_eq_landau_n_rt_iii1_t19 s j) (l_e_st_eq_landau_n_rt_iii1_t31 s s1)))))))). Time Defined. (* constant 2011 *) Definition l_e_st_eq_landau_n_rt_rp_more : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lrt ksi x) (l_e_st_eq_landau_n_rt_urt eta x)))). Time Defined. (* constant 2012 *) Definition l_e_st_eq_landau_n_rt_rp_iii2_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (u:l_e_st_eq_landau_n_rt_urt eta x), p)))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta x0)), l_e_st_eq_landau_n_rt_lrt ksi x0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (u:l_e_st_eq_landau_n_rt_urt eta x), p)))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta x0)) => l_ande1 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta x0) a))))))). Time Defined. (* constant 2013 *) Definition l_e_st_eq_landau_n_rt_rp_iii2_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (u:l_e_st_eq_landau_n_rt_urt eta x), p)))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta x0)), l_e_st_eq_landau_n_rt_urt eta x0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (u:l_e_st_eq_landau_n_rt_urt eta x), p)))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta x0)) => l_ande2 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta x0) a))))))). Time Defined. (* constant 2014 *) Definition l_e_st_eq_landau_n_rt_rp_iii2_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (u:l_e_st_eq_landau_n_rt_urt eta x), p)))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta x0)), p))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (u:l_e_st_eq_landau_n_rt_urt eta x), p)))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta x0)) => p1 x0 (l_e_st_eq_landau_n_rt_rp_iii2_t1 ksi eta m p p1 x0 a) (l_e_st_eq_landau_n_rt_rp_iii2_t2 ksi eta m p p1 x0 a)))))))). Time Defined. (* constant 2015 *) Definition l_e_st_eq_landau_n_rt_rp_moreapp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (u:l_e_st_eq_landau_n_rt_urt eta x), p)))), p))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (u:l_e_st_eq_landau_n_rt_urt eta x), p)))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lrt ksi x) (l_e_st_eq_landau_n_rt_urt eta x)) m p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_and (l_e_st_eq_landau_n_rt_lrt ksi x) (l_e_st_eq_landau_n_rt_urt eta x)) => l_e_st_eq_landau_n_rt_rp_iii2_t3 ksi eta m p p1 x t))))))). Time Defined. (* constant 2016 *) Definition l_e_st_eq_landau_n_rt_rp_less : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_urt ksi x) (l_e_st_eq_landau_n_rt_lrt eta x)))). Time Defined. (* constant 2017 *) Definition l_e_st_eq_landau_n_rt_rp_iii2_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (u:l_e_st_eq_landau_n_rt_lrt eta x), p)))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta x0)), l_e_st_eq_landau_n_rt_urt ksi x0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (u:l_e_st_eq_landau_n_rt_lrt eta x), p)))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta x0)) => l_ande1 (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta x0) a))))))). Time Defined. (* constant 2018 *) Definition l_e_st_eq_landau_n_rt_rp_iii2_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (u:l_e_st_eq_landau_n_rt_lrt eta x), p)))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta x0)), l_e_st_eq_landau_n_rt_lrt eta x0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (u:l_e_st_eq_landau_n_rt_lrt eta x), p)))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta x0)) => l_ande2 (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta x0) a))))))). Time Defined. (* constant 2019 *) Definition l_e_st_eq_landau_n_rt_rp_iii2_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (u:l_e_st_eq_landau_n_rt_lrt eta x), p)))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta x0)), p))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (u:l_e_st_eq_landau_n_rt_lrt eta x), p)))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta x0)) => p1 x0 (l_e_st_eq_landau_n_rt_rp_iii2_t4 ksi eta l p p1 x0 a) (l_e_st_eq_landau_n_rt_rp_iii2_t5 ksi eta l p p1 x0 a)))))))). Time Defined. (* constant 2020 *) Definition l_e_st_eq_landau_n_rt_rp_lessapp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (u:l_e_st_eq_landau_n_rt_lrt eta x), p)))), p))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (u:l_e_st_eq_landau_n_rt_lrt eta x), p)))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_urt ksi x) (l_e_st_eq_landau_n_rt_lrt eta x)) l p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_and (l_e_st_eq_landau_n_rt_urt ksi x) (l_e_st_eq_landau_n_rt_lrt eta x)) => l_e_st_eq_landau_n_rt_rp_iii2_t6 ksi eta l p p1 x t))))))). Time Defined. (* constant 2021 *) Definition l_e_st_eq_landau_n_rt_rp_2121_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (ux:l_e_st_eq_landau_n_rt_urt eta x0), l_and (l_e_st_eq_landau_n_rt_urt eta x0) (l_e_st_eq_landau_n_rt_lrt ksi x0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt eta x0) => l_andi (l_e_st_eq_landau_n_rt_urt eta x0) (l_e_st_eq_landau_n_rt_lrt ksi x0) ux lx)))))). Time Defined. (* constant 2022 *) Definition l_e_st_eq_landau_n_rt_rp_2121_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (ux:l_e_st_eq_landau_n_rt_urt eta x0), l_e_st_eq_landau_n_rt_rp_less eta ksi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt eta x0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_urt eta x) (l_e_st_eq_landau_n_rt_lrt ksi x)) x0 (l_e_st_eq_landau_n_rt_rp_2121_t1 ksi eta m x0 lx ux))))))). Time Defined. (* constant 2023 *) Definition l_e_st_eq_landau_n_rt_rp_satz121 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_less eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_moreapp ksi eta m (l_e_st_eq_landau_n_rt_rp_less eta ksi) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_urt eta x) => l_e_st_eq_landau_n_rt_rp_2121_t2 ksi eta m x t u)))))). Time Defined. (* constant 2024 *) Definition l_e_st_eq_landau_n_rt_rp_2122_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), l_and (l_e_st_eq_landau_n_rt_lrt eta x0) (l_e_st_eq_landau_n_rt_urt ksi x0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => l_andi (l_e_st_eq_landau_n_rt_lrt eta x0) (l_e_st_eq_landau_n_rt_urt ksi x0) lx ux)))))). Time Defined. (* constant 2025 *) Definition l_e_st_eq_landau_n_rt_rp_2122_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), l_e_st_eq_landau_n_rt_rp_more eta ksi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lrt eta x) (l_e_st_eq_landau_n_rt_urt ksi x)) x0 (l_e_st_eq_landau_n_rt_rp_2122_t1 ksi eta l x0 ux lx))))))). Time Defined. (* constant 2026 *) Definition l_e_st_eq_landau_n_rt_rp_satz122 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_more eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_lessapp ksi eta l (l_e_st_eq_landau_n_rt_rp_more eta ksi) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta x) => l_e_st_eq_landau_n_rt_rp_2122_t2 ksi eta l x t u)))))). Time Defined. (* constant 2027 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (ux:l_e_st_eq_landau_n_rt_urt eta x0), l_not (l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt eta x0) => l_e_st_isset_th3 l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta) x0 lx ux)))))). Time Defined. (* constant 2028 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (ux:l_e_st_eq_landau_n_rt_urt eta x0), l_not (l_e_st_eq_landau_n_rt_rp_is ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt eta x0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)) (l_e_st_eq_landau_n_rt_rp_2123_t1 ksi eta m x0 lx ux) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_ise ksi eta t))))))). Time Defined. (* constant 2029 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_not (l_e_st_eq_landau_n_rt_rp_is ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_moreapp ksi eta m (l_not (l_e_st_eq_landau_n_rt_rp_is ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_urt eta x) => l_e_st_eq_landau_n_rt_rp_2123_t2 ksi eta m x t u)))))). Time Defined. (* constant 2030 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_ec (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_ec_th2 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_2123_t3 ksi eta t))). Time Defined. (* constant 2031 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), l_not (l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => l_e_st_isset_th4 l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_lcl eta) (l_e_st_eq_landau_n_rt_lcl ksi) x0 lx ux)))))). Time Defined. (* constant 2032 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), l_not (l_e_st_eq_landau_n_rt_rp_is ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)) (l_e_st_eq_landau_n_rt_rp_2123_t5 ksi eta l x0 ux lx) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_ise ksi eta t))))))). Time Defined. (* constant 2033 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_not (l_e_st_eq_landau_n_rt_rp_is ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_lessapp ksi eta l (l_not (l_e_st_eq_landau_n_rt_rp_is ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta x) => l_e_st_eq_landau_n_rt_rp_2123_t6 ksi eta l x t u)))))). Time Defined. (* constant 2034 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_ec (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_ec_th1 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (fun (t:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_2123_t7 ksi eta t))). Time Defined. (* constant 2035 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (ux:l_e_st_eq_landau_n_rt_urt eta x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), l_e_st_eq_landau_n_rt_less x0 y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt eta x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => l_e_st_eq_landau_n_rt_cutapp2a ksi x0 lx y0 uy)))))))))). Time Defined. (* constant 2036 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (ux:l_e_st_eq_landau_n_rt_urt eta x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), l_e_st_eq_landau_n_rt_more x0 y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt eta x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => l_e_st_eq_landau_n_rt_cutapp2b eta y0 ly x0 ux)))))))))). Time Defined. (* constant 2037 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (ux:l_e_st_eq_landau_n_rt_urt eta x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), l_con)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt eta x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => l_mp (l_e_st_eq_landau_n_rt_less x0 y0) l_con (l_e_st_eq_landau_n_rt_rp_2123_t9 ksi eta m l x0 lx ux y0 uy ly) (l_ec3e23 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_satz81b x0 y0) (l_e_st_eq_landau_n_rt_rp_2123_t10 ksi eta m l x0 lx ux y0 uy ly)))))))))))). Time Defined. (* constant 2038 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (ux:l_e_st_eq_landau_n_rt_urt eta x0), l_con))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt eta x0) => l_e_st_eq_landau_n_rt_rp_lessapp ksi eta l l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta x) => l_e_st_eq_landau_n_rt_rp_2123_t11 ksi eta m l x0 lx ux x t u)))))))))). Time Defined. (* constant 2039 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_con)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_moreapp ksi eta m l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_urt eta x) => l_e_st_eq_landau_n_rt_rp_2123_t12 ksi eta m l x t u))))))). Time Defined. (* constant 2040 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (t:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_2123_t13 ksi eta m t)))). Time Defined. (* constant 2041 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_ec (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_ec_th1 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_2123_t14 ksi eta t))). Time Defined. (* constant 2042 *) Definition l_e_st_eq_landau_n_rt_rp_2123_a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_ec3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_ec3_th6 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_2123_t4 ksi eta) (l_e_st_eq_landau_n_rt_rp_2123_t15 ksi eta) (l_e_st_eq_landau_n_rt_rp_2123_t8 ksi eta))). Time Defined. (* constant 2043 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_or3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_or3i1 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) i))). Time Defined. (* constant 2044 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_nis ksi eta), l_not (l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_nis ksi eta) => l_imp_th3 (l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)) (l_e_st_eq_landau_n_rt_rp_is ksi eta) n (fun (t:l_e_is (l_e_st_set l_e_st_eq_landau_n_rt_rat) (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta)) => l_e_st_eq_landau_n_rt_rp_isi ksi eta t)))). Time Defined. (* constant 2045 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_nis ksi eta), l_or (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_more eta ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_nis ksi eta) => l_e_st_isset_th5 l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_lcl ksi) (l_e_st_eq_landau_n_rt_lcl eta) (l_e_st_eq_landau_n_rt_rp_2123_t17 ksi eta n)))). Time Defined. (* constant 2046 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_nis ksi eta), l_or (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_nis ksi eta) => l_or_th8 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_more eta ksi) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_2123_t18 ksi eta n) (fun (t:l_e_st_eq_landau_n_rt_rp_more eta ksi) => l_e_st_eq_landau_n_rt_rp_satz121 eta ksi t)))). Time Defined. (* constant 2047 *) Definition l_e_st_eq_landau_n_rt_rp_2123_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_nis ksi eta), l_or3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_nis ksi eta) => l_or3_th7 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_2123_t19 ksi eta n)))). Time Defined. (* constant 2048 *) Definition l_e_st_eq_landau_n_rt_rp_2123_b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_or3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_or3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta)) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_2123_t16 ksi eta t) (fun (t:l_e_st_eq_landau_n_rt_rp_nis ksi eta) => l_e_st_eq_landau_n_rt_rp_2123_t20 ksi eta t))). Time Defined. (* constant 2049 *) Definition l_e_st_eq_landau_n_rt_rp_satz123 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_orec3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_orec3i (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_2123_b ksi eta) (l_e_st_eq_landau_n_rt_rp_2123_a ksi eta))). Time Defined. (* constant 2050 *) Definition l_e_st_eq_landau_n_rt_rp_satz123a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_or3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_orec3e1 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123 ksi eta))). Time Defined. (* constant 2051 *) Definition l_e_st_eq_landau_n_rt_rp_satz123b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_ec3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_orec3e2 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123 ksi eta))). Time Defined. (* constant 2052 *) Definition l_e_st_eq_landau_n_rt_rp_moreis : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta))). Time Defined. (* constant 2053 *) Definition l_e_st_eq_landau_n_rt_rp_lessis : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta))). Time Defined. (* constant 2054 *) Definition l_e_st_eq_landau_n_rt_rp_satz124 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), l_e_st_eq_landau_n_rt_rp_lessis eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_less eta ksi) (l_e_st_eq_landau_n_rt_rp_is eta ksi) m (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz121 ksi eta t) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_symis l_e_st_eq_landau_n_rt_cut ksi eta t)))). Time Defined. (* constant 2055 *) Definition l_e_st_eq_landau_n_rt_rp_satz125 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), l_e_st_eq_landau_n_rt_rp_moreis eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more eta ksi) (l_e_st_eq_landau_n_rt_rp_is eta ksi) l (fun (t:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz122 ksi eta t) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_symis l_e_st_eq_landau_n_rt_cut ksi eta t)))). Time Defined. (* constant 2056 *) Definition l_e_st_eq_landau_n_rt_rp_ismore1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta), l_e_st_eq_landau_n_rt_rp_more eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_more u zeta) ksi eta m i))))). Time Defined. (* constant 2057 *) Definition l_e_st_eq_landau_n_rt_rp_ismore2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more zeta ksi), l_e_st_eq_landau_n_rt_rp_more zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more zeta ksi) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_more zeta u) ksi eta m i))))). Time Defined. (* constant 2058 *) Definition l_e_st_eq_landau_n_rt_rp_isless1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi zeta), l_e_st_eq_landau_n_rt_rp_less eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi zeta) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_less u zeta) ksi eta l i))))). Time Defined. (* constant 2059 *) Definition l_e_st_eq_landau_n_rt_rp_isless2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less zeta ksi), l_e_st_eq_landau_n_rt_rp_less zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less zeta ksi) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_less zeta u) ksi eta l i))))). Time Defined. (* constant 2060 *) Definition l_e_st_eq_landau_n_rt_rp_ismoreis1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi zeta), l_e_st_eq_landau_n_rt_rp_moreis eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi zeta) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_moreis u zeta) ksi eta m i))))). Time Defined. (* constant 2061 *) Definition l_e_st_eq_landau_n_rt_rp_ismoreis2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis zeta ksi), l_e_st_eq_landau_n_rt_rp_moreis zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis zeta ksi) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_moreis zeta u) ksi eta m i))))). Time Defined. (* constant 2062 *) Definition l_e_st_eq_landau_n_rt_rp_islessis1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi zeta), l_e_st_eq_landau_n_rt_rp_lessis eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi zeta) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_lessis u zeta) ksi eta l i))))). Time Defined. (* constant 2063 *) Definition l_e_st_eq_landau_n_rt_rp_islessis2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis zeta ksi), l_e_st_eq_landau_n_rt_rp_lessis zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis zeta ksi) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_lessis zeta u) ksi eta l i))))). Time Defined. (* constant 2064 *) Definition l_e_st_eq_landau_n_rt_rp_moreisi2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_moreis ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_ori2 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) i))). Time Defined. (* constant 2065 *) Definition l_e_st_eq_landau_n_rt_rp_lessisi2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_lessis ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_ori2 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) i))). Time Defined. (* constant 2066 *) Definition l_e_st_eq_landau_n_rt_rp_moreisi1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_moreis ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_ori1 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) m))). Time Defined. (* constant 2067 *) Definition l_e_st_eq_landau_n_rt_rp_lessisi1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_lessis ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_ori1 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) l))). Time Defined. (* constant 2068 *) Definition l_e_st_eq_landau_n_rt_rp_ismore12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta), l_e_st_eq_landau_n_rt_rp_more eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta) => l_e_st_eq_landau_n_rt_rp_ismore2 zeta upsilon eta j (l_e_st_eq_landau_n_rt_rp_ismore1 ksi eta zeta i m)))))))). Time Defined. (* constant 2069 *) Definition l_e_st_eq_landau_n_rt_rp_isless12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi zeta), l_e_st_eq_landau_n_rt_rp_less eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi zeta) => l_e_st_eq_landau_n_rt_rp_isless2 zeta upsilon eta j (l_e_st_eq_landau_n_rt_rp_isless1 ksi eta zeta i l)))))))). Time Defined. (* constant 2070 *) Definition l_e_st_eq_landau_n_rt_rp_ismoreis12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi zeta), l_e_st_eq_landau_n_rt_rp_moreis eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi zeta) => l_e_st_eq_landau_n_rt_rp_ismoreis2 zeta upsilon eta j (l_e_st_eq_landau_n_rt_rp_ismoreis1 ksi eta zeta i m)))))))). Time Defined. (* constant 2071 *) Definition l_e_st_eq_landau_n_rt_rp_islessis12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi zeta), l_e_st_eq_landau_n_rt_rp_lessis eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi zeta) => l_e_st_eq_landau_n_rt_rp_islessis2 zeta upsilon eta j (l_e_st_eq_landau_n_rt_rp_islessis1 ksi eta zeta i l)))))))). Time Defined. (* constant 2072 *) Definition l_e_st_eq_landau_n_rt_rp_satz123c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => l_ec3_th7 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123b ksi eta) (l_comor (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) m)))). Time Defined. (* constant 2073 *) Definition l_e_st_eq_landau_n_rt_rp_satz123d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => l_ec3_th9 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123b ksi eta) l))). Time Defined. (* constant 2074 *) Definition l_e_st_eq_landau_n_rt_rp_satz123e : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)), l_e_st_eq_landau_n_rt_rp_lessis ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)) => l_or3_th2 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123a ksi eta) n))). Time Defined. (* constant 2075 *) Definition l_e_st_eq_landau_n_rt_rp_satz123f : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)), l_e_st_eq_landau_n_rt_rp_moreis ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)) => l_comor (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_or3_th3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123a ksi eta) n)))). Time Defined. (* constant 2076 *) Definition l_e_st_eq_landau_n_rt_rp_satz123g : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_not (l_e_st_eq_landau_n_rt_rp_lessis ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_or_th3 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_ec3e23 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123b ksi eta) m) (l_ec3e21 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123b ksi eta) m)))). Time Defined. (* constant 2077 *) Definition l_e_st_eq_landau_n_rt_rp_satz123h : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_not (l_e_st_eq_landau_n_rt_rp_moreis ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_or_th3 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_ec3e32 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123b ksi eta) l) (l_ec3e31 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123b ksi eta) l)))). Time Defined. (* constant 2078 *) Definition l_e_st_eq_landau_n_rt_rp_satz123j : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_moreis ksi eta)), l_e_st_eq_landau_n_rt_rp_less ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_moreis ksi eta)) => l_or3e3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123a ksi eta) (l_or_th5 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) n) (l_or_th4 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) n)))). Time Defined. (* constant 2079 *) Definition l_e_st_eq_landau_n_rt_rp_satz123k : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_lessis ksi eta)), l_e_st_eq_landau_n_rt_rp_more ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_lessis ksi eta)) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123a ksi eta) (l_or_th4 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) n) (l_or_th5 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) n)))). Time Defined. (* constant 2080 *) Definition l_e_st_eq_landau_n_rt_rp_2126_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less eta zeta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (ly:l_e_st_eq_landau_n_rt_lrt zeta y0), l_e_st_eq_landau_n_rt_less x0 y0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less eta zeta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (ly:l_e_st_eq_landau_n_rt_lrt zeta y0) => l_e_st_eq_landau_n_rt_cutapp2a eta x0 lx y0 uy))))))))))). Time Defined. (* constant 2081 *) Definition l_e_st_eq_landau_n_rt_rp_2126_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less eta zeta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (ly:l_e_st_eq_landau_n_rt_lrt zeta y0), l_e_st_eq_landau_n_rt_urt ksi y0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less eta zeta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (ly:l_e_st_eq_landau_n_rt_lrt zeta y0) => l_e_st_eq_landau_n_rt_rp_satz119a ksi x0 ux y0 (l_e_st_eq_landau_n_rt_rp_2126_t1 ksi eta zeta l k x0 ux lx y0 uy ly)))))))))))). Time Defined. (* constant 2082 *) Definition l_e_st_eq_landau_n_rt_rp_2126_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less eta zeta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (ly:l_e_st_eq_landau_n_rt_lrt zeta y0), l_and (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_lrt zeta y0)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less eta zeta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (ly:l_e_st_eq_landau_n_rt_lrt zeta y0) => l_andi (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_lrt zeta y0) (l_e_st_eq_landau_n_rt_rp_2126_t2 ksi eta zeta l k x0 ux lx y0 uy ly) ly))))))))))). Time Defined. (* constant 2083 *) Definition l_e_st_eq_landau_n_rt_rp_2126_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less eta zeta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (ly:l_e_st_eq_landau_n_rt_lrt zeta y0), l_e_st_eq_landau_n_rt_rp_less ksi zeta))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less eta zeta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (ly:l_e_st_eq_landau_n_rt_lrt zeta y0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_urt ksi x) (l_e_st_eq_landau_n_rt_lrt zeta x)) y0 (l_e_st_eq_landau_n_rt_rp_2126_t3 ksi eta zeta l k x0 ux lx y0 uy ly)))))))))))). Time Defined. (* constant 2084 *) Definition l_e_st_eq_landau_n_rt_rp_2126_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less eta zeta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), l_e_st_eq_landau_n_rt_rp_less ksi zeta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less eta zeta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => l_e_st_eq_landau_n_rt_rp_lessapp eta zeta k (l_e_st_eq_landau_n_rt_rp_less ksi zeta) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt eta x) => (fun (u:l_e_st_eq_landau_n_rt_lrt zeta x) => l_e_st_eq_landau_n_rt_rp_2126_t4 ksi eta zeta l k x0 ux lx x t u))))))))))). Time Defined. (* constant 2085 *) Definition l_e_st_eq_landau_n_rt_rp_satz126 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less eta zeta), l_e_st_eq_landau_n_rt_rp_less ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less eta zeta) => l_e_st_eq_landau_n_rt_rp_lessapp ksi eta l (l_e_st_eq_landau_n_rt_rp_less ksi zeta) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta x) => l_e_st_eq_landau_n_rt_rp_2126_t5 ksi eta zeta l k x t u)))))))). Time Defined. (* constant 2086 *) Definition l_e_st_eq_landau_n_rt_rp_trless : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less eta zeta), l_e_st_eq_landau_n_rt_rp_less ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less eta zeta) => l_e_st_eq_landau_n_rt_rp_satz126 ksi eta zeta l k))))). Time Defined. (* constant 2087 *) Definition l_e_st_eq_landau_n_rt_rp_trmore : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more eta zeta), l_e_st_eq_landau_n_rt_rp_more ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more eta zeta) => l_e_st_eq_landau_n_rt_rp_satz122 zeta ksi (l_e_st_eq_landau_n_rt_rp_satz126 zeta eta ksi (l_e_st_eq_landau_n_rt_rp_satz121 eta zeta n) (l_e_st_eq_landau_n_rt_rp_satz121 ksi eta m))))))). Time Defined. (* constant 2088 *) Definition l_e_st_eq_landau_n_rt_rp_satz127a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less eta zeta), l_e_st_eq_landau_n_rt_rp_less ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less eta zeta) => l_orapp (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi zeta) l (fun (u:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_trless ksi eta zeta u k) (fun (u:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_isless1 eta ksi zeta (l_e_symis l_e_st_eq_landau_n_rt_cut ksi eta u) k)))))). Time Defined. (* constant 2089 *) Definition l_e_st_eq_landau_n_rt_rp_satz127b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta), l_e_st_eq_landau_n_rt_rp_less ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta) => l_orapp (l_e_st_eq_landau_n_rt_rp_less eta zeta) (l_e_st_eq_landau_n_rt_rp_is eta zeta) (l_e_st_eq_landau_n_rt_rp_less ksi zeta) k (fun (u:l_e_st_eq_landau_n_rt_rp_less eta zeta) => l_e_st_eq_landau_n_rt_rp_trless ksi eta zeta l u) (fun (u:l_e_st_eq_landau_n_rt_rp_is eta zeta) => l_e_st_eq_landau_n_rt_rp_isless2 eta zeta ksi u l)))))). Time Defined. (* constant 2090 *) Definition l_e_st_eq_landau_n_rt_rp_satz127c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more eta zeta), l_e_st_eq_landau_n_rt_rp_more ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more eta zeta) => l_e_st_eq_landau_n_rt_rp_satz122 zeta ksi (l_e_st_eq_landau_n_rt_rp_satz127b zeta eta ksi (l_e_st_eq_landau_n_rt_rp_satz121 eta zeta n) (l_e_st_eq_landau_n_rt_rp_satz124 ksi eta m))))))). Time Defined. (* constant 2091 *) Definition l_e_st_eq_landau_n_rt_rp_satz127d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis eta zeta), l_e_st_eq_landau_n_rt_rp_more ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis eta zeta) => l_e_st_eq_landau_n_rt_rp_satz122 zeta ksi (l_e_st_eq_landau_n_rt_rp_satz127a zeta eta ksi (l_e_st_eq_landau_n_rt_rp_satz124 eta zeta n) (l_e_st_eq_landau_n_rt_rp_satz121 ksi eta m))))))). Time Defined. (* constant 2092 *) Definition l_e_st_eq_landau_n_rt_rp_2128_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is eta zeta), l_e_st_eq_landau_n_rt_rp_lessis ksi zeta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is eta zeta) => l_e_st_eq_landau_n_rt_rp_lessisi2 ksi zeta (l_e_tris l_e_st_eq_landau_n_rt_cut ksi eta zeta i j)))))))). Time Defined. (* constant 2093 *) Definition l_e_st_eq_landau_n_rt_rp_2128_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_less eta zeta), l_e_st_eq_landau_n_rt_rp_lessis ksi zeta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_less eta zeta) => l_e_st_eq_landau_n_rt_rp_lessisi1 ksi zeta (l_e_st_eq_landau_n_rt_rp_satz127a ksi eta zeta l j)))))))). Time Defined. (* constant 2094 *) Definition l_e_st_eq_landau_n_rt_rp_2128_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_lessis ksi zeta)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_orapp (l_e_st_eq_landau_n_rt_rp_less eta zeta) (l_e_st_eq_landau_n_rt_rp_is eta zeta) (l_e_st_eq_landau_n_rt_rp_lessis ksi zeta) k (fun (t:l_e_st_eq_landau_n_rt_rp_less eta zeta) => l_e_st_eq_landau_n_rt_rp_2128_t2 ksi eta zeta l k i t) (fun (t:l_e_st_eq_landau_n_rt_rp_is eta zeta) => l_e_st_eq_landau_n_rt_rp_2128_t1 ksi eta zeta l k i t))))))). Time Defined. (* constant 2095 *) Definition l_e_st_eq_landau_n_rt_rp_2128_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta), (forall (j:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_lessis ksi zeta)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta) => (fun (j:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_lessisi1 ksi zeta (l_e_st_eq_landau_n_rt_rp_satz127b ksi eta zeta j k))))))). Time Defined. (* constant 2096 *) Definition l_e_st_eq_landau_n_rt_rp_satz128 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta), l_e_st_eq_landau_n_rt_rp_lessis ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta) => l_orapp (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_lessis ksi zeta) l (fun (t:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_2128_t4 ksi eta zeta l k t) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_2128_t3 ksi eta zeta l k t)))))). Time Defined. (* constant 2097 *) Definition l_e_st_eq_landau_n_rt_rp_trlessis : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta), l_e_st_eq_landau_n_rt_rp_lessis ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis eta zeta) => l_e_st_eq_landau_n_rt_rp_satz128 ksi eta zeta l k))))). Time Defined. (* constant 2098 *) Definition l_e_st_eq_landau_n_rt_rp_trmoreis : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis eta zeta), l_e_st_eq_landau_n_rt_rp_moreis ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis eta zeta) => l_e_st_eq_landau_n_rt_rp_satz125 zeta ksi (l_e_st_eq_landau_n_rt_rp_satz128 zeta eta ksi (l_e_st_eq_landau_n_rt_rp_satz124 eta zeta n) (l_e_st_eq_landau_n_rt_rp_satz124 ksi eta m))))))). Time Defined. (* constant 2099 *) Definition l_e_st_eq_landau_n_rt_rp_sumprop1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_and3 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0))))))). Time Defined. (* constant 2100 *) Definition l_e_st_eq_landau_n_rt_rp_sumprop : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), Prop))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x y))))). Time Defined. (* constant 2101 *) Definition l_e_st_eq_landau_n_rt_rp_sum : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_set l_e_st_eq_landau_n_rt_rat)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop ksi eta z))). Time Defined. (* constant 2102 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_and3i (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) lx ly i)))))))). Time Defined. (* constant 2103 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y) y0 (l_e_st_eq_landau_n_rt_rp_iii3_t1 ksi eta z0 x0 lx y0 ly i))))))))). Time Defined. (* constant 2104 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_rp_sumprop ksi eta z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x y)) x0 (l_e_st_eq_landau_n_rt_rp_iii3_t2 ksi eta z0 x0 lx y0 ly i))))))))). Time Defined. (* constant 2105 *) Definition l_e_st_eq_landau_n_rt_rp_sum1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop ksi eta z) z0 (l_e_st_eq_landau_n_rt_rp_iii3_t3 ksi eta z0 x0 lx y0 ly i))))))))). Time Defined. (* constant 2106 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), l_e_st_eq_landau_n_rt_rp_sumprop ksi eta z0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop ksi eta z) z0 i)))))). Time Defined. (* constant 2107 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_lrt ksi x0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0) => l_and3e1 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) py)))))))))). Time Defined. (* constant 2108 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_lrt eta y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0) => l_and3e2 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) py)))))))))). Time Defined. (* constant 2109 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0) => l_and3e3 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) py)))))))))). Time Defined. (* constant 2110 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0), p)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y0) => p1 x0 (l_e_st_eq_landau_n_rt_rp_iii3_t5 ksi eta z0 i p p1 x0 px y0 py) y0 (l_e_st_eq_landau_n_rt_rp_iii3_t6 ksi eta z0 i p p1 x0 px y0 py) (l_e_st_eq_landau_n_rt_rp_iii3_t7 ksi eta z0 i p p1 x0 px y0 py))))))))))). Time Defined. (* constant 2111 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)), p)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y)) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y) px p (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x0 y) => l_e_st_eq_landau_n_rt_rp_iii3_t8 ksi eta z0 i p p1 x0 px y t)))))))))). Time Defined. (* constant 2112 *) Definition l_e_st_eq_landau_n_rt_rp_sumapp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), p)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x y)) (l_e_st_eq_landau_n_rt_rp_iii3_t4 ksi eta z0 i p p1) p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_sumprop1 ksi eta z0 x y)) => l_e_st_eq_landau_n_rt_rp_iii3_t9 ksi eta z0 i p p1 x t)))))))). Time Defined. (* constant 2113 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_less x0 x1))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_cutapp2a ksi x0 lx x1 ux))))))))))))). Time Defined. (* constant 2114 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_less y0 y1))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_cutapp2a eta y0 ly y1 uy))))))))))))). Time Defined. (* constant 2115 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_less z0 (l_e_st_eq_landau_n_rt_pl x1 y1)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_isless1 (l_e_st_eq_landau_n_rt_pl x0 y0) z0 (l_e_st_eq_landau_n_rt_pl x1 y1) (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_pl x0 y0) j) (l_e_st_eq_landau_n_rt_satz98a x0 x1 y0 y1 (l_e_st_eq_landau_n_rt_rp_3129_t1 ksi eta x1 ux y1 uy z0 i x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_rp_3129_t2 ksi eta x1 ux y1 uy z0 i x0 lx y0 ly j))))))))))))))). Time Defined. (* constant 2116 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_nis z0 (l_e_st_eq_landau_n_rt_pl x1 y1)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_ec3e31 (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x1 y1)) (l_e_st_eq_landau_n_rt_more z0 (l_e_st_eq_landau_n_rt_pl x1 y1)) (l_e_st_eq_landau_n_rt_less z0 (l_e_st_eq_landau_n_rt_pl x1 y1)) (l_e_st_eq_landau_n_rt_satz81b z0 (l_e_st_eq_landau_n_rt_pl x1 y1)) (l_e_st_eq_landau_n_rt_rp_3129_t3 ksi eta x1 ux y1 uy z0 i x0 lx y0 ly j)))))))))))))). Time Defined. (* constant 2117 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), l_e_st_eq_landau_n_rt_nis z0 (l_e_st_eq_landau_n_rt_pl x1 y1))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => l_e_st_eq_landau_n_rt_rp_sumapp ksi eta z0 i (l_e_st_eq_landau_n_rt_nis z0 (l_e_st_eq_landau_n_rt_pl x1 y1)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3129_t4 ksi eta x1 ux y1 uy z0 i x t y u v))))))))))))). Time Defined. (* constant 2118 *) Definition l_e_st_eq_landau_n_rt_rp_satz129a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), l_not (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x1 y1) (l_e_st_eq_landau_n_rt_rp_sum ksi eta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => l_imp_th3 (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x1 y1) (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_nis (l_e_st_eq_landau_n_rt_pl x1 y1) (l_e_st_eq_landau_n_rt_pl x1 y1)) (l_weli (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl x1 y1) (l_e_st_eq_landau_n_rt_pl x1 y1)) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl x1 y1))) (fun (t:l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x1 y1) (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => l_e_st_eq_landau_n_rt_rp_3129_t5 ksi eta x1 ux y1 uy (l_e_st_eq_landau_n_rt_pl x1 y1) t))))))). Time Defined. (* constant 2119 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_less z0 (l_e_st_eq_landau_n_rt_pl x0 y0)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_isless2 u0 (l_e_st_eq_landau_n_rt_pl x0 y0) z0 j l))))))))))). Time Defined. (* constant 2120 *) Definition l_e_st_eq_landau_n_rt_rp_3129_z1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_rat))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_ov z0 (l_e_st_eq_landau_n_rt_pl x0 y0)))))))))))). Time Defined. (* constant 2121 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_pl x0 y0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_isless12 z0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_satz110f z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_example1d (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_rp_3129_t6 ksi eta u0 i z0 l x0 lx y0 ly j)))))))))))). Time Defined. (* constant 2122 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j) l_e_st_eq_landau_n_rt_1rt))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_satz106c (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_rp_3129_t7 ksi eta u0 i z0 l x0 lx y0 ly j)))))))))))). Time Defined. (* constant 2123 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) x0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt) x0 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_example1a x0) (l_e_st_eq_landau_n_rt_satz105f (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j) l_e_st_eq_landau_n_rt_1rt x0 (l_e_st_eq_landau_n_rt_rp_3129_t8 ksi eta u0 i z0 l x0 lx y0 ly j))))))))))))). Time Defined. (* constant 2124 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) y0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_ts y0 l_e_st_eq_landau_n_rt_1rt) y0 (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_example1a y0) (l_e_st_eq_landau_n_rt_satz105f (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j) l_e_st_eq_landau_n_rt_1rt y0 (l_e_st_eq_landau_n_rt_rp_3129_t8 ksi eta u0 i z0 l x0 lx y0 ly j))))))))))))). Time Defined. (* constant 2125 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_rp_satz120 ksi x0 lx (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_rp_3129_t9 ksi eta u0 i z0 l x0 lx y0 ly j)))))))))))). Time Defined. (* constant 2126 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_lrt eta (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_rp_satz120 eta y0 ly (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_rp_3129_t10 ksi eta u0 i z0 l x0 lx y0 ly j)))))))))))). Time Defined. (* constant 2127 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j))) z0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j))) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) z0 (l_e_st_eq_landau_n_rt_distpt1 x0 y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_satz110c z0 (l_e_st_eq_landau_n_rt_pl x0 y0))))))))))))). Time Defined. (* constant 2128 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j))) z0 (l_e_st_eq_landau_n_rt_rp_3129_t13 ksi eta u0 i z0 l x0 lx y0 ly j)))))))))))). Time Defined. (* constant 2129 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_rp_sum1 ksi eta z0 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_rp_3129_t11 ksi eta u0 i z0 l x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_3129_z1 ksi eta u0 i z0 l x0 lx y0 ly j)) (l_e_st_eq_landau_n_rt_rp_3129_t12 ksi eta u0 i z0 l x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_rp_3129_t14 ksi eta u0 i z0 l x0 lx y0 ly j)))))))))))). Time Defined. (* constant 2130 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => l_e_st_eq_landau_n_rt_rp_sumapp ksi eta u0 i (l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3129_t15 ksi eta u0 i z0 l x t y u v))))))))))). Time Defined. (* constant 2131 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x1 y0) (l_e_st_eq_landau_n_rt_rp_sum ksi eta))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_rp_sum1 ksi eta (l_e_st_eq_landau_n_rt_pl x1 y0) x1 lx1 y0 ly (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl x1 y0)))))))))))))). Time Defined. (* constant 2132 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x1 y0) (l_e_st_eq_landau_n_rt_pl x0 y0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_satz96a x1 x0 y0 (l_e_st_eq_landau_n_rt_satz83 x0 x1 l))))))))))))). Time Defined. (* constant 2133 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x1 y0) z0)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_ismore2 (l_e_st_eq_landau_n_rt_pl x0 y0) z0 (l_e_st_eq_landau_n_rt_pl x1 y0) (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_pl x0 y0) j) (l_e_st_eq_landau_n_rt_rp_3129_t18 ksi eta z0 i x0 lx y0 ly j x1 lx1 l))))))))))))). Time Defined. (* constant 2134 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_and (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x1 y0) (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x1 y0) z0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_andi (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x1 y0) (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x1 y0) z0) (l_e_st_eq_landau_n_rt_rp_3129_t17 ksi eta z0 i x0 lx y0 ly j x1 lx1 l) (l_e_st_eq_landau_n_rt_rp_3129_t19 ksi eta z0 i x0 lx y0 ly j x1 lx1 l))))))))))))). Time Defined. (* constant 2135 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_more y z0)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_more y z0)) (l_e_st_eq_landau_n_rt_pl x1 y0) (l_e_st_eq_landau_n_rt_rp_3129_t20 ksi eta z0 i x0 lx y0 ly j x1 lx1 l))))))))))))). Time Defined. (* constant 2136 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_more y z0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_cutapp3 ksi x0 lx (l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_more y z0))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_less x0 x) => l_e_st_eq_landau_n_rt_rp_3129_t21 ksi eta z0 i x0 lx y0 ly j x t u)))))))))))). Time Defined. (* constant 2137 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_more y z0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => l_e_st_eq_landau_n_rt_rp_sumapp ksi eta z0 i (l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_more y z0))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3129_t22 ksi eta z0 i x t y u v))))))))). Time Defined. (* constant 2138 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t24 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => l_e_st_eq_landau_n_rt_cut2 (l_e_st_eq_landau_n_rt_rp_sum ksi eta) (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_rp_sum1 ksi eta (l_e_st_eq_landau_n_rt_pl x0 y0) x0 lx y0 ly (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl x0 y0))) (l_e_st_eq_landau_n_rt_pl x1 y1) (l_e_st_eq_landau_n_rt_rp_satz129a ksi eta x1 ux y1 uy) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_less y x) => l_e_st_eq_landau_n_rt_rp_3129_t16 ksi eta x t y u)))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) => l_e_st_eq_landau_n_rt_rp_3129_t23 ksi eta x t)))))))))))). Time Defined. (* constant 2139 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t25 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => l_e_st_eq_landau_n_rt_cutapp1b eta (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt eta y) => l_e_st_eq_landau_n_rt_rp_3129_t24 ksi eta x0 lx y0 ly x1 ux y t)))))))))). Time Defined. (* constant 2140 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t26 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => l_e_st_eq_landau_n_rt_cutapp1b ksi (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => l_e_st_eq_landau_n_rt_rp_3129_t25 ksi eta x0 lx y0 ly x t)))))))). Time Defined. (* constant 2141 *) Definition l_e_st_eq_landau_n_rt_rp_3129_t27 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_eq_landau_n_rt_cutapp1a eta (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta y) => l_e_st_eq_landau_n_rt_rp_3129_t26 ksi eta x0 lx y t)))))). Time Defined. (* constant 2142 *) Definition l_e_st_eq_landau_n_rt_rp_satz129 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_cutapp1a ksi (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => l_e_st_eq_landau_n_rt_rp_3129_t27 ksi eta x t)))). Time Defined. (* constant 2143 *) Definition l_e_st_eq_landau_n_rt_rp_pl : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_cutof (l_e_st_eq_landau_n_rt_rp_sum ksi eta) (l_e_st_eq_landau_n_rt_rp_satz129 ksi eta))). Time Defined. (* constant 2144 *) Definition l_e_st_eq_landau_n_rt_rp_lrtpl : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_rp_ine (l_e_st_eq_landau_n_rt_rp_sum ksi eta) (l_e_st_eq_landau_n_rt_rp_satz129 ksi eta) z0 (l_e_st_eq_landau_n_rt_rp_sum1 ksi eta z0 x0 lx y0 ly i))))))))). Time Defined. (* constant 2145 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_not (l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_isp1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_not (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_sum ksi eta))) (l_e_st_eq_landau_n_rt_pl x0 y0) z0 (l_e_st_eq_landau_n_rt_rp_satz129a ksi eta x0 ux y0 uy) i)))))))). Time Defined. (* constant 2146 *) Definition l_e_st_eq_landau_n_rt_rp_urtpl : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_imp_th3 (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0) (l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta)) (l_e_st_eq_landau_n_rt_rp_iii3_t10 ksi eta z0 x0 ux y0 uy i) (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0) => l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_rp_sum ksi eta) (l_e_st_eq_landau_n_rt_rp_satz129 ksi eta) z0 t))))))))). Time Defined. (* constant 2147 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_sum ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_rp_sum ksi eta) (l_e_st_eq_landau_n_rt_rp_satz129 ksi eta) z0 lz)))))). Time Defined. (* constant 2148 *) Definition l_e_st_eq_landau_n_rt_rp_plapp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))), p)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)), p)))))) => l_e_st_eq_landau_n_rt_rp_sumapp ksi eta z0 (l_e_st_eq_landau_n_rt_rp_iii3_t11 ksi eta z0 lz p p1) p p1)))))). Time Defined. (* constant 2149 *) Definition l_e_st_eq_landau_n_rt_rp_ispl1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pl u zeta) ksi eta i)))). Time Defined. (* constant 2150 *) Definition l_e_st_eq_landau_n_rt_rp_ispl2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pl zeta u) ksi eta i)))). Time Defined. (* constant 2151 *) Definition l_e_st_eq_landau_n_rt_rp_ispl12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_ispl1 ksi eta zeta i) (l_e_st_eq_landau_n_rt_rp_ispl2 zeta upsilon eta j))))))). Time Defined. (* constant 2152 *) Definition l_e_st_eq_landau_n_rt_rp_3130_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y0 x0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_tris l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_pl y0 x0) i (l_e_st_eq_landau_n_rt_compl x0 y0)))))))))). Time Defined. (* constant 2153 *) Definition l_e_st_eq_landau_n_rt_rp_3130_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta ksi) z0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_rp_lrtpl eta ksi z0 y0 ly x0 lx (l_e_st_eq_landau_n_rt_rp_3130_t1 ksi eta z0 lz x0 lx y0 ly i)))))))))). Time Defined. (* constant 2154 *) Definition l_e_st_eq_landau_n_rt_rp_3130_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta ksi) z0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) z0) => l_e_st_eq_landau_n_rt_rp_plapp ksi eta z0 lz (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta ksi) z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3130_t2 ksi eta z0 lz x t y u v))))))))). Time Defined. (* constant 2155 *) Definition l_e_st_eq_landau_n_rt_rp_satz130 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi eta) (l_e_st_eq_landau_n_rt_rp_pl eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_pl ksi eta) (l_e_st_eq_landau_n_rt_rp_pl eta ksi) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) x) => l_e_st_eq_landau_n_rt_rp_3130_t3 ksi eta x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta ksi) x) => l_e_st_eq_landau_n_rt_rp_3130_t3 eta ksi x t)))). Time Defined. (* constant 2156 *) Definition l_e_st_eq_landau_n_rt_rp_compl : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi eta) (l_e_st_eq_landau_n_rt_rp_pl eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz130 ksi eta)). Time Defined. (* constant 2157 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) v0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 z0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) v0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 z0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_pl v0 z0) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) i (l_e_st_eq_landau_n_rt_ispl1 v0 (l_e_st_eq_landau_n_rt_pl x0 y0) z0 j) (l_e_st_eq_landau_n_rt_asspl1 x0 y0 z0)))))))))))))))). Time Defined. (* constant 2158 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) v0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 z0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_pl y0 z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) v0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 z0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_rp_lrtpl eta zeta (l_e_st_eq_landau_n_rt_pl y0 z0) y0 ly z0 lz (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))))))))). Time Defined. (* constant 2159 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) v0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 z0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) v0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 z0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x0 y0)) => l_e_st_eq_landau_n_rt_rp_lrtpl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta) u0 x0 lx (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_rp_3131_t2 ksi eta zeta u0 lu v0 lv z0 lz i x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_rp_3131_t1 ksi eta zeta u0 lu v0 lv z0 lz i x0 lx y0 ly j)))))))))))))))). Time Defined. (* constant 2160 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) v0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) v0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 z0)) => l_e_st_eq_landau_n_rt_rp_plapp ksi eta v0 lv (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3131_t3 ksi eta zeta u0 lu v0 lv z0 lz i x t y u v))))))))))))))). Time Defined. (* constant 2161 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0) => l_e_st_eq_landau_n_rt_rp_plapp (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta u0 lu (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt zeta y) => (fun (v:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3131_t4 ksi eta zeta u0 lu x t y u v)))))))))). Time Defined. (* constant 2162 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_pl x0 v0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl x0 y0) z0) i (l_e_st_eq_landau_n_rt_ispl2 v0 (l_e_st_eq_landau_n_rt_pl y0 z0) x0 j) (l_e_st_eq_landau_n_rt_asspl2 x0 y0 z0)))))))))))))))). Time Defined. (* constant 2163 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) (l_e_st_eq_landau_n_rt_pl x0 y0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_rp_lrtpl ksi eta (l_e_st_eq_landau_n_rt_pl x0 y0) x0 lx y0 ly (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl x0 y0))))))))))))))))). Time Defined. (* constant 2164 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_rp_lrtpl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta u0 (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_rp_3131_t7 ksi eta zeta u0 lu x0 lx v0 lv i y0 ly z0 lz j) z0 lz (l_e_st_eq_landau_n_rt_rp_3131_t6 ksi eta zeta u0 lu x0 lx v0 lv i y0 ly z0 lz j)))))))))))))))). Time Defined. (* constant 2165 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 v0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 v0)) => l_e_st_eq_landau_n_rt_rp_plapp eta zeta v0 lv (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt zeta y) => (fun (v:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3131_t8 ksi eta zeta u0 lu x0 lx v0 lv i x t y u v))))))))))))))). Time Defined. (* constant 2166 *) Definition l_e_st_eq_landau_n_rt_rp_3131_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => l_e_st_eq_landau_n_rt_rp_plapp ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta) u0 lu (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) y) => (fun (v:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3131_t9 ksi eta zeta u0 lu x t y u v)))))))))). Time Defined. (* constant 2167 *) Definition l_e_st_eq_landau_n_rt_rp_satz131 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) x) => l_e_st_eq_landau_n_rt_rp_3131_t5 ksi eta zeta x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) x) => l_e_st_eq_landau_n_rt_rp_3131_t10 ksi eta zeta x t))))). Time Defined. (* constant 2168 *) Definition l_e_st_eq_landau_n_rt_rp_asspl1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz131 ksi eta zeta))). Time Defined. (* constant 2169 *) Definition l_e_st_eq_landau_n_rt_rp_asspl2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_satz131 ksi eta zeta)))). Time Defined. (* constant 2170 *) Definition l_e_st_eq_landau_n_rt_rp_3132_prop1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt ksi y0))))). Time Defined. (* constant 2171 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_more y0 x0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0) => l_e_st_eq_landau_n_rt_cutapp2b ksi x0 (l_ande1 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt ksi y0) p) y0 (l_ande2 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt ksi y0) p)))))). Time Defined. (* constant 2172 *) Definition l_e_st_eq_landau_n_rt_rp_3132_prop2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0), Prop))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0) => l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t1 ksi a0 x0 y0 p)) a0))))). Time Defined. (* constant 2173 *) Definition l_e_st_eq_landau_n_rt_rp_3132_prop3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0) (forall (t:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_rp_3132_prop2 ksi a0 x0 y0 t))))). Time Defined. (* constant 2174 *) Definition l_e_st_eq_landau_n_rt_rp_3132_prop4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x y)))). Time Defined. (* constant 2175 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_rt_more y0 x0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx y0 uy)))))). Time Defined. (* constant 2176 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0)) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))) => l_e_st_eq_landau_n_rt_satz96d (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy)) x0 m)))))))). Time Defined. (* constant 2177 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0)) y0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))) => l_e_st_eq_landau_n_rt_ismore2 (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))) y0 (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0)) (l_e_st_eq_landau_n_rt_satz101c y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy)) (l_e_st_eq_landau_n_rt_rp_3132_t3 ksi a0 x0 lx y0 uy n m))))))))). Time Defined. (* constant 2178 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))), l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))) => l_e_st_eq_landau_n_rt_rp_satz119 ksi y0 uy (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0)) (l_e_st_eq_landau_n_rt_rp_3132_t4 ksi a0 x0 lx y0 uy n m))))))))). Time Defined. (* constant 2179 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))), l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn n) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))) => l_somei l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) n (l_e_st_eq_landau_n_rt_rp_3132_t5 ksi a0 x0 lx y0 uy n m))))))))). Time Defined. (* constant 2180 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_someapp l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))) (l_e_st_eq_landau_n_rt_satz115 a0 (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))) (l_e_st_eq_landau_n_some (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0)))) (fun (x:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t2 ksi a0 x0 lx y0 uy))) => l_e_st_eq_landau_n_rt_rp_3132_t6 ksi a0 x0 lx y0 uy x t)))))))). Time Defined. (* constant 2181 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), l_e_st_eq_landau_n_lb (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => l_ande1 (l_e_st_eq_landau_n_lb (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) (l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0))) m)))))))). Time Defined. (* constant 2182 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => l_ande2 (l_e_st_eq_landau_n_lb (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) (l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0))) m)))))))). Time Defined. (* constant 2183 *) Definition l_e_st_eq_landau_n_rt_rp_3132_u0 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rat))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_pl x0 a0))))))))). Time Defined. (* constant 2184 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0) a0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt a0) (l_e_st_eq_landau_n_rt_ts a0 l_e_st_eq_landau_n_rt_1rt) a0 (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_rtofn u) l_e_st_eq_landau_n_rt_1rt a0 (l_e_st_eq_landau_n_rt_isnert u l_e_st_eq_landau_n_1 i)) (l_e_st_eq_landau_n_rt_comts l_e_st_eq_landau_n_rt_1rt a0) (l_e_st_eq_landau_n_rt_example1a a0)))))))))). Time Defined. (* constant 2185 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 x)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0) a0 (l_e_st_eq_landau_n_rt_rp_3132_t9 ksi a0 x0 lx y0 uy u m) (l_e_st_eq_landau_n_rt_rp_3132_t10 ksi a0 x0 lx y0 uy u m i)))))))))). Time Defined. (* constant 2186 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_andi (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)) lx (l_e_st_eq_landau_n_rt_rp_3132_t11 ksi a0 x0 lx y0 uy u m i)))))))))). Time Defined. (* constant 2187 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), (forall (p:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)), l_e_st_eq_landau_n_rt_rp_3132_prop2 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i) p)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => (fun (p:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)) => l_e_symis l_e_st_eq_landau_n_rt_rat a0 (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i) x0 (l_e_st_eq_landau_n_rt_rp_3132_t1 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i) p)) (l_e_st_eq_landau_n_rt_satz101g (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i) x0 a0 (l_e_st_eq_landau_n_rt_rp_3132_t1 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i) p) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i))))))))))))). Time Defined. (* constant 2188 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_andi (l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)) (forall (t:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)), l_e_st_eq_landau_n_rt_rp_3132_prop2 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i) t) (l_e_st_eq_landau_n_rt_rp_3132_t12 ksi a0 x0 lx y0 uy u m i) (fun (t:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i)) => l_e_st_eq_landau_n_rt_rp_3132_t13 ksi a0 x0 lx y0 uy u m i t)))))))))). Time Defined. (* constant 2189 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y) (l_e_st_eq_landau_n_rt_rp_3132_u0 ksi a0 x0 lx y0 uy u m i) (l_e_st_eq_landau_n_rt_rp_3132_t14 ksi a0 x0 lx y0 uy u m i)))))))))). Time Defined. (* constant 2190 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (i:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x y)) x0 (l_e_st_eq_landau_n_rt_rp_3132_t15 ksi a0 x0 lx y0 uy u m i)))))))))). Time Defined. (* constant 2191 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn u) l_e_st_eq_landau_n_rt_1rt))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_rtofn u) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_fr u l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr u l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_satz111d u l_e_st_eq_landau_n_1 o)))))))))). Time Defined. (* constant 2192 *) Definition l_e_st_eq_landau_n_rt_rp_3132_um10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rat))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rtofn u) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_3132_t17 ksi a0 x0 lx y0 uy u m o)))))))))). Time Defined. (* constant 2193 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_satz112g (l_e_st_eq_landau_n_rt_rtofn u) (l_e_st_eq_landau_n_rt_natrti u) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_natrti l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_3132_t17 ksi a0 x0 lx y0 uy u m o)))))))))). Time Defined. (* constant 2194 *) Definition l_e_st_eq_landau_n_rt_rp_3132_um1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_nat))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t18 ksi a0 x0 lx y0 uy u m o)))))))))). Time Defined. (* constant 2195 *) Definition l_e_st_eq_landau_n_rt_rp_3132_v0 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rat))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) a0)))))))))). Time Defined. (* constant 2196 *) Definition l_e_st_eq_landau_n_rt_rp_3132_w0 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rat))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0)))))))))). Time Defined. (* constant 2197 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rtofn u)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) l_e_st_eq_landau_n_rt_1rt) (l_e_st_eq_landau_n_rt_rtofn u) (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_satz101e (l_e_st_eq_landau_n_rt_rtofn u) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_3132_t17 ksi a0 x0 lx y0 uy u m o)) (l_e_st_eq_landau_n_rt_satz94a (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) l_e_st_eq_landau_n_rt_1rt)))))))))). Time Defined. (* constant 2198 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_lessf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr u l_e_st_eq_landau_n_1)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_lesse (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rtofn u) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr u l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclassn (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t18 ksi a0 x0 lx y0 uy u m o)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr u l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_3132_t19 ksi a0 x0 lx y0 uy u m o)))))))))). Time Defined. (* constant 2199 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o) u))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_satz111c (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o) u (l_e_st_eq_landau_n_rt_rp_3132_t20 ksi a0 x0 lx y0 uy u m o)))))))))). Time Defined. (* constant 2200 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_not (l_e_st_eq_landau_n_lessis u (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_imp_th3 (l_e_st_eq_landau_n_lessis u (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o)) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o) u) (l_e_st_eq_landau_n_satz10h (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o) u (l_e_st_eq_landau_n_rt_rp_3132_t21 ksi a0 x0 lx y0 uy u m o)) (fun (t:l_e_st_eq_landau_n_lessis u (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o)) => l_e_st_eq_landau_n_satz14 u (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o) t)))))))))). Time Defined. (* constant 2201 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o)) a0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_et (l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o)) a0))) (l_imp_th3 (l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o)) a0))) (l_e_st_eq_landau_n_lessis u (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o)) (l_e_st_eq_landau_n_rt_rp_3132_t22 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t8 ksi a0 x0 lx y0 uy u m (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o)))))))))))). Time Defined. (* constant 2202 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t24 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_isp1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts x a0))) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_3132_um1 ksi a0 x0 lx y0 uy u m o)) (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t23 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_isrtn1 (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t18 ksi a0 x0 lx y0 uy u m o))))))))))). Time Defined. (* constant 2203 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t25 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_andi (l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o)) (l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)) (l_e_st_eq_landau_n_rt_rp_3132_t24 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t9 ksi a0 x0 lx y0 uy u m)))))))))). Time Defined. (* constant 2204 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t26 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) a0) a0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) a0) a0) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) a0) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt a0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) l_e_st_eq_landau_n_rt_1rt) a0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0) (l_e_st_eq_landau_n_rt_ispl2 a0 (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt a0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) a0) (l_e_st_eq_landau_n_rt_example1d a0)) (l_e_st_eq_landau_n_rt_distpt1 (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) l_e_st_eq_landau_n_rt_1rt a0) (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) l_e_st_eq_landau_n_rt_1rt) (l_e_st_eq_landau_n_rt_rtofn u) a0 (l_e_st_eq_landau_n_rt_satz101e (l_e_st_eq_landau_n_rt_rtofn u) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_3132_t17 ksi a0 x0 lx y0 uy u m o)))))))))))). Time Defined. (* constant 2205 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t27 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) a0) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) a0) (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) a0) a0)) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_asspl1 x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) a0) a0) (l_e_st_eq_landau_n_rt_ispl2 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_3132_um10 ksi a0 x0 lx y0 uy u m o) a0) a0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn u) a0) x0 (l_e_st_eq_landau_n_rt_rp_3132_t26 ksi a0 x0 lx y0 uy u m o))))))))))). Time Defined. (* constant 2206 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t28 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), (forall (p:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)), l_e_st_eq_landau_n_rt_rp_3132_prop2 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o) p)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => (fun (p:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)) => l_e_symis l_e_st_eq_landau_n_rt_rat a0 (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t1 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o) p)) (l_e_st_eq_landau_n_rt_satz101g (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) a0 (l_e_st_eq_landau_n_rt_rp_3132_t1 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o) p) (l_e_st_eq_landau_n_rt_rp_3132_t27 ksi a0 x0 lx y0 uy u m o)))))))))))). Time Defined. (* constant 2207 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t29 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_andi (l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)) (forall (t:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)), l_e_st_eq_landau_n_rt_rp_3132_prop2 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o) t) (l_e_st_eq_landau_n_rt_rp_3132_t25 ksi a0 x0 lx y0 uy u m o) (fun (t:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o)) => l_e_st_eq_landau_n_rt_rp_3132_t28 ksi a0 x0 lx y0 uy u m o t)))))))))). Time Defined. (* constant 2208 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t30 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) y)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) y) (l_e_st_eq_landau_n_rt_rp_3132_w0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t29 ksi a0 x0 lx y0 uy u m o)))))))))). Time Defined. (* constant 2209 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t31 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), (forall (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => (fun (o:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x y)) (l_e_st_eq_landau_n_rt_rp_3132_v0 ksi a0 x0 lx y0 uy u m o) (l_e_st_eq_landau_n_rt_rp_3132_t30 ksi a0 x0 lx y0 uy u m o)))))))))). Time Defined. (* constant 2210 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t32 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (u:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u), l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_min (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) a0))) u) => l_orapp (l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0) (l_e_st_eq_landau_n_satz24 u) (fun (t:l_e_st_eq_landau_n_more u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_3132_t31 ksi a0 x0 lx y0 uy u m t) (fun (t:l_e_st_eq_landau_n_is u l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_3132_t16 ksi a0 x0 lx y0 uy u m t))))))))). Time Defined. (* constant 2211 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t34 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_someapp l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn y) a0))) x) (l_e_st_eq_landau_n_satz27 (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn y) a0))) (l_e_st_eq_landau_n_rt_rp_3132_t7 ksi a0 x0 lx y0 uy)) (l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0) (fun (x:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_min (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn y) a0))) x) => l_e_st_eq_landau_n_rt_rp_3132_t32 ksi a0 x0 lx y0 uy x t)))))))). Time Defined. (* constant 2212 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t35 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_eq_landau_n_rt_cutapp1b ksi (l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi y) => l_e_st_eq_landau_n_rt_rp_3132_t34 ksi a0 x0 lx y t)))))). Time Defined. (* constant 2213 *) Definition l_e_st_eq_landau_n_rt_rp_satz132 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (a0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_and (l_e_st_eq_landau_n_rt_lrt ksi x) (l_e_st_eq_landau_n_rt_urt ksi y)) (forall (t:l_and (l_e_st_eq_landau_n_rt_lrt ksi x) (l_e_st_eq_landau_n_rt_urt ksi y)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x (l_ande1 (l_e_st_eq_landau_n_rt_lrt ksi x) (l_e_st_eq_landau_n_rt_urt ksi y) t) y (l_ande2 (l_e_st_eq_landau_n_rt_lrt ksi x) (l_e_st_eq_landau_n_rt_urt ksi y) t))) a0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_cutapp1a ksi (l_e_st_eq_landau_n_rt_rp_3132_prop4 ksi a0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => l_e_st_eq_landau_n_rt_rp_3132_t35 ksi a0 x t)))). Time Defined. (* constant 2214 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t36 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0) => l_ande1 (l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0) (forall (t:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_rp_3132_prop2 ksi a0 x0 y0 t) p3)))))))). Time Defined. (* constant 2215 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t37 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_rp_3132_prop2 ksi a0 x0 y0 (l_e_st_eq_landau_n_rt_rp_3132_t36 ksi p a0 p1 x0 s y0 p3))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0) => l_r_ande2 (l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0) (fun (t:l_e_st_eq_landau_n_rt_rp_3132_prop1 ksi a0 x0 y0) => l_e_st_eq_landau_n_rt_rp_3132_prop2 ksi a0 x0 y0 t) p3)))))))). Time Defined. (* constant 2216 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t38 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_lrt ksi x0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0) => l_ande1 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_rp_3132_t36 ksi p a0 p1 x0 s y0 p3))))))))). Time Defined. (* constant 2217 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t39 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_urt ksi y0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0) => l_ande2 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_rp_3132_t36 ksi p a0 p1 x0 s y0 p3))))))))). Time Defined. (* constant 2218 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t40 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 (l_e_st_eq_landau_n_rt_rp_3132_t38 ksi p a0 p1 x0 s y0 p3) y0 (l_e_st_eq_landau_n_rt_rp_3132_t39 ksi p a0 p1 x0 s y0 p3))) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t1 ksi a0 x0 y0 (l_e_st_eq_landau_n_rt_rp_3132_t36 ksi p a0 p1 x0 s y0 p3))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0) => l_e_st_eq_landau_n_rt_satz101g y0 x0 (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 (l_e_st_eq_landau_n_rt_rp_3132_t38 ksi p a0 p1 x0 s y0 p3) y0 (l_e_st_eq_landau_n_rt_rp_3132_t39 ksi p a0 p1 x0 s y0 p3))) (l_e_st_eq_landau_n_rt_rp_3132_t1 ksi a0 x0 y0 (l_e_st_eq_landau_n_rt_rp_3132_t36 ksi p a0 p1 x0 s y0 p3)) (l_e_st_eq_landau_n_rt_satz101c y0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 (l_e_st_eq_landau_n_rt_rp_3132_t38 ksi p a0 p1 x0 s y0 p3) y0 (l_e_st_eq_landau_n_rt_rp_3132_t39 ksi p a0 p1 x0 s y0 p3))))))))))). Time Defined. (* constant 2219 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t41 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 (l_e_st_eq_landau_n_rt_rp_3132_t38 ksi p a0 p1 x0 s y0 p3) y0 (l_e_st_eq_landau_n_rt_rp_3132_t39 ksi p a0 p1 x0 s y0 p3))) a0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 (l_e_st_eq_landau_n_rt_rp_3132_t38 ksi p a0 p1 x0 s y0 p3) y0 (l_e_st_eq_landau_n_rt_rp_3132_t39 ksi p a0 p1 x0 s y0 p3))) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_3132_t1 ksi a0 x0 y0 (l_e_st_eq_landau_n_rt_rp_3132_t36 ksi p a0 p1 x0 s y0 p3))) a0 (l_e_st_eq_landau_n_rt_rp_3132_t40 ksi p a0 p1 x0 s y0 p3) (l_e_st_eq_landau_n_rt_rp_3132_t37 ksi p a0 p1 x0 s y0 p3))))))))). Time Defined. (* constant 2220 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t42 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0), p)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (p3:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y0) => p1 x0 (l_e_st_eq_landau_n_rt_rp_3132_t38 ksi p a0 p1 x0 s y0 p3) y0 (l_e_st_eq_landau_n_rt_rp_3132_t39 ksi p a0 p1 x0 s y0 p3) (l_e_st_eq_landau_n_rt_rp_3132_t41 ksi p a0 p1 x0 s y0 p3))))))))). Time Defined. (* constant 2221 *) Definition l_e_st_eq_landau_n_rt_rp_3132_t43 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)), p)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y)) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y) s p (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x0 y) => l_e_st_eq_landau_n_rt_rp_3132_t42 ksi p a0 p1 x0 s y t)))))))). Time Defined. (* constant 2222 *) Definition l_e_st_eq_landau_n_rt_rp_satz132app : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (p:Prop), (forall (a0:l_e_st_eq_landau_n_rt_rat), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))), p)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (p:Prop) => (fun (a0:l_e_st_eq_landau_n_rt_rat) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) a0), p)))))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x y)) (l_e_st_eq_landau_n_rt_rp_satz132 ksi a0) p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_3132_prop3 ksi a0 x y)) => l_e_st_eq_landau_n_rt_rp_3132_t43 ksi p a0 p1 x t)))))). Time Defined. (* constant 2223 *) Definition l_e_st_eq_landau_n_rt_rp_3133_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt ksi u0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x0 y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt ksi u0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0) => l_e_tris l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_pl x0 (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu))) (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_satz101d u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) (l_e_st_eq_landau_n_rt_ispl2 (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0 x0 i)))))))))). Time Defined. (* constant 2224 *) Definition l_e_st_eq_landau_n_rt_rp_3133_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt ksi u0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) u0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt ksi u0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0) => l_e_st_eq_landau_n_rt_rp_lrtpl ksi eta u0 x0 lx y0 ly (l_e_st_eq_landau_n_rt_rp_3133_t1 ksi eta y0 ly x0 lx u0 uu i)))))))))). Time Defined. (* constant 2225 *) Definition l_e_st_eq_landau_n_rt_rp_3133_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt ksi u0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0), l_and (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) u0) (l_e_st_eq_landau_n_rt_urt ksi u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt ksi u0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0) => l_andi (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) u0) (l_e_st_eq_landau_n_rt_urt ksi u0) (l_e_st_eq_landau_n_rt_rp_3133_t2 ksi eta y0 ly x0 lx u0 uu i) uu))))))))). Time Defined. (* constant 2226 *) Definition l_e_st_eq_landau_n_rt_rp_3133_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt ksi u0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi eta) ksi))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt ksi u0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn u0 x0 (l_e_st_eq_landau_n_rt_cutapp2b ksi x0 lx u0 uu)) y0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi eta) x) (l_e_st_eq_landau_n_rt_urt ksi x)) u0 (l_e_st_eq_landau_n_rt_rp_3133_t3 ksi eta y0 ly x0 lx u0 uu i)))))))))). Time Defined. (* constant 2227 *) Definition l_e_st_eq_landau_n_rt_rp_3133_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi eta) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => l_e_st_eq_landau_n_rt_rp_satz132app ksi (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi eta) ksi) y0 (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt ksi y) => (fun (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) y0) => l_e_st_eq_landau_n_rt_rp_3133_t4 ksi eta y0 ly x t y u v))))))))). Time Defined. (* constant 2228 *) Definition l_e_st_eq_landau_n_rt_rp_satz133 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi eta) ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_cutapp1a eta (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi eta) ksi) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta x) => l_e_st_eq_landau_n_rt_rp_3133_t5 ksi eta x t)))). Time Defined. (* constant 2229 *) Definition l_e_st_eq_landau_n_rt_rp_satz133a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_pl ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_pl ksi eta) ksi (l_e_st_eq_landau_n_rt_rp_satz133 ksi eta))). Time Defined. (* constant 2230 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_urt eta x0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_rp_satz119a eta y0 uy x0 l)))))))))). Time Defined. (* constant 2231 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_more x0 y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_satz83 y0 x0 l)))))))))). Time Defined. (* constant 2232 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt zeta u0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (uz:l_e_st_eq_landau_n_rt_urt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l)) u0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt zeta u0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (uz:l_e_st_eq_landau_n_rt_urt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) => l_e_tris l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) u0) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l)) u0) (l_e_st_eq_landau_n_rt_satz101f z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_ispl1 (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l)) u0 i)))))))))))))))). Time Defined. (* constant 2233 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt zeta u0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (uz:l_e_st_eq_landau_n_rt_urt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_pl x0 u0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt zeta u0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (uz:l_e_st_eq_landau_n_rt_urt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l)) u0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) u0) (l_e_st_eq_landau_n_rt_pl x0 u0) (l_e_st_eq_landau_n_rt_ispl2 z0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l)) u0) y0 (l_e_st_eq_landau_n_rt_rp_3134_t3 ksi eta zeta m y0 ly uy x0 lx l u0 lu z0 uz i)) (l_e_st_eq_landau_n_rt_asspl2 y0 (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l)) u0) (l_e_st_eq_landau_n_rt_ispl1 (l_e_st_eq_landau_n_rt_pl y0 (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) x0 u0 (l_e_st_eq_landau_n_rt_satz101c x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l)))))))))))))))))). Time Defined. (* constant 2234 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt zeta u0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (uz:l_e_st_eq_landau_n_rt_urt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_pl y0 z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt zeta u0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (uz:l_e_st_eq_landau_n_rt_urt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_rp_lrtpl ksi zeta (l_e_st_eq_landau_n_rt_pl y0 z0) x0 lx u0 lu (l_e_st_eq_landau_n_rt_rp_3134_t4 ksi eta zeta m y0 ly uy x0 lx l u0 lu z0 uz i)))))))))))))))). Time Defined. (* constant 2235 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt zeta u0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (uz:l_e_st_eq_landau_n_rt_urt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_pl y0 z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt zeta u0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (uz:l_e_st_eq_landau_n_rt_urt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_rp_urtpl eta zeta (l_e_st_eq_landau_n_rt_pl y0 z0) y0 uy z0 uz (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))))))))). Time Defined. (* constant 2236 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt zeta u0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (uz:l_e_st_eq_landau_n_rt_urt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))), l_and (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt zeta u0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (uz:l_e_st_eq_landau_n_rt_urt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) => l_andi (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_rp_3134_t5 ksi eta zeta m y0 ly uy x0 lx l u0 lu z0 uz i) (l_e_st_eq_landau_n_rt_rp_3134_t6 ksi eta zeta m y0 ly uy x0 lx l u0 lu z0 uz i)))))))))))))))). Time Defined. (* constant 2237 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt zeta u0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (uz:l_e_st_eq_landau_n_rt_urt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt zeta u0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (uz:l_e_st_eq_landau_n_rt_urt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn z0 u0 (l_e_st_eq_landau_n_rt_cutapp2b zeta u0 lu z0 uz)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) x) (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) x)) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_rp_3134_t7 ksi eta zeta m y0 ly uy x0 lx l u0 lu z0 uz i)))))))))))))))). Time Defined. (* constant 2238 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_rp_satz132app zeta (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt zeta x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt zeta y) => (fun (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b zeta x t y u)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta zeta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_rp_3134_t8 ksi eta zeta m y0 ly uy x0 lx l x t y u v))))))))))))))). Time Defined. (* constant 2239 *) Definition l_e_st_eq_landau_n_rt_rp_3134_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => l_e_st_eq_landau_n_rt_cutapp3 ksi y0 ly (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_less y0 x) => l_e_st_eq_landau_n_rt_rp_3134_t9 ksi eta zeta m y0 ly uy x t u)))))))))). Time Defined. (* constant 2240 *) Definition l_e_st_eq_landau_n_rt_rp_satz134 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_moreapp ksi eta m (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_urt eta x) => l_e_st_eq_landau_n_rt_rp_3134_t10 ksi eta zeta m x t u))))))). Time Defined. (* constant 2241 *) Definition l_e_st_eq_landau_n_rt_rp_satz135a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz134 ksi eta zeta m)))). Time Defined. (* constant 2242 *) Definition l_e_st_eq_landau_n_rt_rp_satz135b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_ispl1 ksi eta zeta i)))). Time Defined. (* constant 2243 *) Definition l_e_st_eq_landau_n_rt_rp_satz135c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz134 eta ksi zeta (l_e_st_eq_landau_n_rt_rp_satz122 ksi eta l)))))). Time Defined. (* constant 2244 *) Definition l_e_st_eq_landau_n_rt_rp_satz135d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta eta) (l_e_st_eq_landau_n_rt_rp_compl ksi zeta) (l_e_st_eq_landau_n_rt_rp_compl eta zeta) (l_e_st_eq_landau_n_rt_rp_satz135a ksi eta zeta m))))). Time Defined. (* constant 2245 *) Definition l_e_st_eq_landau_n_rt_rp_satz135e : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_ispl2 ksi eta zeta i)))). Time Defined. (* constant 2246 *) Definition l_e_st_eq_landau_n_rt_rp_satz135f : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta eta) (l_e_st_eq_landau_n_rt_rp_compl ksi zeta) (l_e_st_eq_landau_n_rt_rp_compl eta zeta) (l_e_st_eq_landau_n_rt_rp_satz135c ksi eta zeta l))))). Time Defined. (* constant 2247 *) Definition l_e_st_eq_landau_n_rt_rp_satz135g : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_ismore2 (l_e_st_eq_landau_n_rt_rp_pl ksi upsilon) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_ispl1 ksi eta upsilon i) (l_e_st_eq_landau_n_rt_rp_satz135d zeta upsilon ksi m))))))). Time Defined. (* constant 2248 *) Definition l_e_st_eq_landau_n_rt_rp_satz135h : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl upsilon eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_pl upsilon eta) (l_e_st_eq_landau_n_rt_rp_compl ksi zeta) (l_e_st_eq_landau_n_rt_rp_compl eta upsilon) (l_e_st_eq_landau_n_rt_rp_satz135g ksi eta zeta upsilon i m))))))). Time Defined. (* constant 2249 *) Definition l_e_st_eq_landau_n_rt_rp_satz135j : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less zeta upsilon) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_pl ksi upsilon) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_ispl1 ksi eta upsilon i) (l_e_st_eq_landau_n_rt_rp_satz135f zeta upsilon ksi l))))))). Time Defined. (* constant 2250 *) Definition l_e_st_eq_landau_n_rt_rp_satz135k : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl upsilon eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less zeta upsilon) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_pl upsilon eta) (l_e_st_eq_landau_n_rt_rp_compl ksi zeta) (l_e_st_eq_landau_n_rt_rp_compl eta upsilon) (l_e_st_eq_landau_n_rt_rp_satz135j ksi eta zeta upsilon i l))))))). Time Defined. (* constant 2251 *) Definition l_e_st_eq_landau_n_rt_rp_3136_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_or3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz123a ksi eta))). Time Defined. (* constant 2252 *) Definition l_e_st_eq_landau_n_rt_rp_3136_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_ec3 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)))). Time Defined. (* constant 2253 *) Definition l_e_st_eq_landau_n_rt_rp_satz136a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)), l_e_st_eq_landau_n_rt_rp_more ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) => l_ec3_th11 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_3136_t1 ksi eta zeta) (l_e_st_eq_landau_n_rt_rp_3136_t2 ksi eta zeta) (fun (u:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135b ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135a ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135c ksi eta zeta u) m)))). Time Defined. (* constant 2254 *) Definition l_e_st_eq_landau_n_rt_rp_satz136b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)), l_e_st_eq_landau_n_rt_rp_is ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) => l_ec3_th10 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_3136_t1 ksi eta zeta) (l_e_st_eq_landau_n_rt_rp_3136_t2 ksi eta zeta) (fun (u:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135b ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135a ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135c ksi eta zeta u) i)))). Time Defined. (* constant 2255 *) Definition l_e_st_eq_landau_n_rt_rp_satz136c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)), l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) => l_ec3_th12 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_3136_t1 ksi eta zeta) (l_e_st_eq_landau_n_rt_rp_3136_t2 ksi eta zeta) (fun (u:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135b ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135a ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135c ksi eta zeta u) l)))). Time Defined. (* constant 2256 *) Definition l_e_st_eq_landau_n_rt_rp_satz136d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta)), l_e_st_eq_landau_n_rt_rp_more ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta)) => l_e_st_eq_landau_n_rt_rp_satz136a ksi eta zeta (l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta eta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_compl zeta ksi) (l_e_st_eq_landau_n_rt_rp_compl zeta eta) m))))). Time Defined. (* constant 2257 *) Definition l_e_st_eq_landau_n_rt_rp_satz136e : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta)), l_e_st_eq_landau_n_rt_rp_is ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta)) => l_e_st_eq_landau_n_rt_rp_satz136b ksi eta zeta (l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_compl ksi zeta) i (l_e_st_eq_landau_n_rt_rp_compl zeta eta)))))). Time Defined. (* constant 2258 *) Definition l_e_st_eq_landau_n_rt_rp_satz136f : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta)), l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl zeta eta)) => l_e_st_eq_landau_n_rt_rp_satz136c ksi eta zeta (l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl zeta ksi) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl zeta eta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_compl zeta ksi) (l_e_st_eq_landau_n_rt_rp_compl zeta eta) l))))). Time Defined. (* constant 2259 *) Definition l_e_st_eq_landau_n_rt_rp_3137_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz134 ksi eta zeta m)))))). Time Defined. (* constant 2260 *) Definition l_e_st_eq_landau_n_rt_rp_3137_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl zeta eta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_pl upsilon eta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_compl zeta eta) (l_e_st_eq_landau_n_rt_rp_compl upsilon eta) (l_e_st_eq_landau_n_rt_rp_satz134 zeta upsilon eta n))))))). Time Defined. (* constant 2261 *) Definition l_e_st_eq_landau_n_rt_rp_satz137 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_trmore (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_3137_t1 ksi eta zeta upsilon m n) (l_e_st_eq_landau_n_rt_rp_3137_t2 ksi eta zeta upsilon m n))))))). Time Defined. (* constant 2262 *) Definition l_e_st_eq_landau_n_rt_rp_satz137a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz137 eta ksi upsilon zeta (l_e_st_eq_landau_n_rt_rp_satz122 ksi eta l) (l_e_st_eq_landau_n_rt_rp_satz122 zeta upsilon k)))))))). Time Defined. (* constant 2263 *) Definition l_e_st_eq_landau_n_rt_rp_satz138a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_orapp (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon)) m (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz137 ksi eta zeta upsilon t n) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_satz135g ksi eta zeta upsilon t n))))))). Time Defined. (* constant 2264 *) Definition l_e_st_eq_landau_n_rt_rp_satz138b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => l_orapp (l_e_st_eq_landau_n_rt_rp_more zeta upsilon) (l_e_st_eq_landau_n_rt_rp_is zeta upsilon) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon)) n (fun (t:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz137 ksi eta zeta upsilon m t) (fun (t:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz135h zeta upsilon ksi eta t m))))))). Time Defined. (* constant 2265 *) Definition l_e_st_eq_landau_n_rt_rp_satz138c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz138a eta ksi upsilon zeta (l_e_st_eq_landau_n_rt_rp_satz125 ksi eta l) (l_e_st_eq_landau_n_rt_rp_satz122 zeta upsilon k)))))))). Time Defined. (* constant 2266 *) Definition l_e_st_eq_landau_n_rt_rp_satz138d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz138b eta ksi upsilon zeta (l_e_st_eq_landau_n_rt_rp_satz122 ksi eta l) (l_e_st_eq_landau_n_rt_rp_satz125 zeta upsilon k)))))))). Time Defined. (* constant 2267 *) Definition l_e_st_eq_landau_n_rt_rp_3139_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_st_eq_landau_n_rt_rp_moreisi2 (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_ispl12 ksi eta zeta upsilon i j))))))))). Time Defined. (* constant 2268 *) Definition l_e_st_eq_landau_n_rt_rp_3139_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (o:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (o:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_moreisi1 (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_satz138a ksi eta zeta upsilon m o))))))))). Time Defined. (* constant 2269 *) Definition l_e_st_eq_landau_n_rt_rp_3139_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_orapp (l_e_st_eq_landau_n_rt_rp_more zeta upsilon) (l_e_st_eq_landau_n_rt_rp_is zeta upsilon) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon)) n (fun (t:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_3139_t2 ksi eta zeta upsilon m n i t) (fun (t:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_st_eq_landau_n_rt_rp_3139_t1 ksi eta zeta upsilon m n i t)))))))). Time Defined. (* constant 2270 *) Definition l_e_st_eq_landau_n_rt_rp_3139_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), (forall (o:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => (fun (o:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_moreisi1 (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_satz138b ksi eta zeta upsilon o n)))))))). Time Defined. (* constant 2271 *) Definition l_e_st_eq_landau_n_rt_rp_satz139 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => l_orapp (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon)) m (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_3139_t4 ksi eta zeta upsilon m n t) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_3139_t3 ksi eta zeta upsilon m n t))))))). Time Defined. (* constant 2272 *) Definition l_e_st_eq_landau_n_rt_rp_satz139a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis zeta upsilon), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz124 (l_e_st_eq_landau_n_rt_rp_pl eta upsilon) (l_e_st_eq_landau_n_rt_rp_pl ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz139 eta ksi upsilon zeta (l_e_st_eq_landau_n_rt_rp_satz125 ksi eta l) (l_e_st_eq_landau_n_rt_rp_satz125 zeta upsilon k)))))))). Time Defined. (* constant 2273 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi), l_e_st_eq_landau_n_rt_rp_more ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi) => l_e_st_eq_landau_n_rt_rp_ismore1 (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi eta i (l_e_st_eq_landau_n_rt_rp_satz133 eta phi)))))). Time Defined. (* constant 2274 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (phi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_satz123d ksi eta l) (fun (t:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi) => l_e_st_eq_landau_n_rt_rp_3140_t1 ksi eta l phi t))))). Time Defined. (* constant 2275 *) Definition l_e_st_eq_landau_n_rt_rp_vorbemerkung140 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), l_not (l_e_st_eq_landau_n_rt_rp_some (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta a) ksi))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => l_some_th5 l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta a) ksi) (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_3140_t2 ksi eta l a)))). Time Defined. (* constant 2276 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more phi psi), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more phi psi) => l_e_st_eq_landau_n_rt_rp_satz135d phi psi eta m))))). Time Defined. (* constant 2277 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more phi psi), l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more phi psi) => l_ec3e21 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_3140_t3 ksi eta phi psi m)))))). Time Defined. (* constant 2278 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less phi psi), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less phi psi) => l_e_st_eq_landau_n_rt_rp_satz135f phi psi eta l))))). Time Defined. (* constant 2279 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less phi psi), l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less phi psi) => l_ec3e31 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_3140_t5 ksi eta phi psi l)))))). Time Defined. (* constant 2280 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_nis phi psi), l_or (l_e_st_eq_landau_n_rt_rp_more phi psi) (l_e_st_eq_landau_n_rt_rp_less phi psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_nis phi psi) => l_or3_th1 (l_e_st_eq_landau_n_rt_rp_is phi psi) (l_e_st_eq_landau_n_rt_rp_more phi psi) (l_e_st_eq_landau_n_rt_rp_less phi psi) (l_e_st_eq_landau_n_rt_rp_satz123a phi psi) n))))). Time Defined. (* constant 2281 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_nis phi psi), l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_nis phi psi) => l_orapp (l_e_st_eq_landau_n_rt_rp_more phi psi) (l_e_st_eq_landau_n_rt_rp_less phi psi) (l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_st_eq_landau_n_rt_rp_3140_t7 ksi eta phi psi n) (fun (t:l_e_st_eq_landau_n_rt_rp_more phi psi) => l_e_st_eq_landau_n_rt_rp_3140_t4 ksi eta phi psi t) (fun (t:l_e_st_eq_landau_n_rt_rp_less phi psi) => l_e_st_eq_landau_n_rt_rp_3140_t6 ksi eta phi psi t)))))). Time Defined. (* constant 2282 *) Definition l_e_st_eq_landau_n_rt_rp_satz140b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta psi) ksi), l_e_st_eq_landau_n_rt_rp_is phi psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta psi) ksi) => l_imp_th7 (l_e_st_eq_landau_n_rt_rp_is phi psi) (l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_weli (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi)) (l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl eta phi) (l_e_st_eq_landau_n_rt_rp_pl eta psi) ksi i j)) (fun (t:l_e_st_eq_landau_n_rt_rp_nis phi psi) => l_e_st_eq_landau_n_rt_rp_3140_t8 ksi eta phi psi t))))))). Time Defined. (* constant 2283 *) Definition l_e_st_eq_landau_n_rt_rp_diffprop1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_more x0 y0) (forall (t:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 t))))))). Time Defined. (* constant 2284 *) Definition l_e_st_eq_landau_n_rt_rp_diffprop2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_and3 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta y0) (l_e_st_eq_landau_n_rt_rp_diffprop1 ksi eta z0 x0 y0)))))). Time Defined. (* constant 2285 *) Definition l_e_st_eq_landau_n_rt_rp_diffprop : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), Prop))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x y))))). Time Defined. (* constant 2286 *) Definition l_e_st_eq_landau_n_rt_rp_diff : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_set l_e_st_eq_landau_n_rt_rat)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop ksi eta z))). Time Defined. (* constant 2287 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t11a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)), (forall (m1:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m1))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) => (fun (m1:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_tris l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m) (l_e_st_eq_landau_n_rt_mn x0 y0 m1) i (l_e_st_eq_landau_n_rt_satz101g x0 y0 (l_e_st_eq_landau_n_rt_mn x0 y0 m) m1 (l_e_st_eq_landau_n_rt_satz101c x0 y0 m)))))))))))). Time Defined. (* constant 2288 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)), l_e_st_eq_landau_n_rt_rp_diffprop1 ksi eta z0 x0 y0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) => l_andi (l_e_st_eq_landau_n_rt_more x0 y0) (forall (t:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 t)) m (fun (t:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_rp_iii3_t11a ksi eta z0 x0 lx y0 uy m i t)))))))))). Time Defined. (* constant 2289 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)), l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) => l_and3i (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta y0) (l_e_st_eq_landau_n_rt_rp_diffprop1 ksi eta z0 x0 y0) lx uy (l_e_st_eq_landau_n_rt_rp_iii3_t12 ksi eta z0 x0 lx y0 uy m i)))))))))). Time Defined. (* constant 2290 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y) y0 (l_e_st_eq_landau_n_rt_rp_iii3_t13 ksi eta z0 x0 lx y0 uy m i)))))))))). Time Defined. (* constant 2291 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)), l_e_st_eq_landau_n_rt_rp_diffprop ksi eta z0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x y)) x0 (l_e_st_eq_landau_n_rt_rp_iii3_t14 ksi eta z0 x0 lx y0 uy m i)))))))))). Time Defined. (* constant 2292 *) Definition l_e_st_eq_landau_n_rt_rp_diff1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 m)) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop ksi eta z) z0 (l_e_st_eq_landau_n_rt_rp_iii3_t15 ksi eta z0 x0 lx y0 uy m i)))))))))). Time Defined. (* constant 2293 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), l_e_st_eq_landau_n_rt_rp_diffprop ksi eta z0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop ksi eta z) z0 i)))))). Time Defined. (* constant 2294 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_lrt ksi x0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0) => l_and3e1 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta y0) (l_e_st_eq_landau_n_rt_rp_diffprop1 ksi eta z0 x0 y0) py)))))))))). Time Defined. (* constant 2295 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_urt eta y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0) => l_and3e2 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta y0) (l_e_st_eq_landau_n_rt_rp_diffprop1 ksi eta z0 x0 y0) py)))))))))). Time Defined. (* constant 2296 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_rp_diffprop1 ksi eta z0 x0 y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0) => l_and3e3 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_urt eta y0) (l_e_st_eq_landau_n_rt_rp_diffprop1 ksi eta z0 x0 y0) py)))))))))). Time Defined. (* constant 2297 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_more x0 y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0) => l_ande1 (l_e_st_eq_landau_n_rt_more x0 y0) (forall (t:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 t)) (l_e_st_eq_landau_n_rt_rp_iii3_t19 ksi eta z0 i p p1 x0 px y0 py))))))))))). Time Defined. (* constant 2298 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_iii3_t20 ksi eta z0 i p p1 x0 px y0 py)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0) => l_r_ande2 (l_e_st_eq_landau_n_rt_more x0 y0) (fun (t:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 t)) (l_e_st_eq_landau_n_rt_rp_iii3_t19 ksi eta z0 i p p1 x0 px y0 py))))))))))). Time Defined. (* constant 2299 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0), p)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y0) => p1 x0 (l_e_st_eq_landau_n_rt_rp_iii3_t17 ksi eta z0 i p p1 x0 px y0 py) y0 (l_e_st_eq_landau_n_rt_rp_iii3_t18 ksi eta z0 i p p1 x0 px y0 py) (l_e_st_eq_landau_n_rt_rp_iii3_t20 ksi eta z0 i p p1 x0 px y0 py) (l_e_st_eq_landau_n_rt_rp_iii3_t21 ksi eta z0 i p p1 x0 px y0 py))))))))))). Time Defined. (* constant 2300 *) Definition l_e_st_eq_landau_n_rt_rp_iii3_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)), p)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y)) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y) px p (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x0 y) => l_e_st_eq_landau_n_rt_rp_iii3_t22 ksi eta z0 i p p1 x0 px y t)))))))))). Time Defined. (* constant 2301 *) Definition l_e_st_eq_landau_n_rt_rp_diffapp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))), p)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt eta y), (forall (v:l_e_st_eq_landau_n_rt_more x y), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)), p))))))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x y)) (l_e_st_eq_landau_n_rt_rp_iii3_t16 ksi eta z0 i p p1) p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_diffprop2 ksi eta z0 x y)) => l_e_st_eq_landau_n_rt_rp_iii3_t23 ksi eta z0 i p p1 x t)))))))). Time Defined. (* constant 2302 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_more x0 y0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_rp_3134_t2 ksi eta eta m y0 ly uy x0 lx l))))))))). Time Defined. (* constant 2303 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t9 ksi eta m y0 ly uy x0 lx l)) (l_e_st_eq_landau_n_rt_rp_diff ksi eta)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_rp_diff1 ksi eta (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t9 ksi eta m y0 ly uy x0 lx l)) x0 lx y0 uy (l_e_st_eq_landau_n_rt_rp_3140_t9 ksi eta m y0 ly uy x0 lx l) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t9 ksi eta m y0 ly uy x0 lx l)))))))))))). Time Defined. (* constant 2304 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_less z0 x0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_mn x0 y0 n) z0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 n) y0) x0 (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n) j) (l_e_st_eq_landau_n_rt_satz101e x0 y0 n) (l_e_st_eq_landau_n_rt_satz94a (l_e_st_eq_landau_n_rt_mn x0 y0 n) y0)))))))))))))). Time Defined. (* constant 2305 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_less z0 x1))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_trless z0 x0 x1 (l_e_st_eq_landau_n_rt_rp_3140_t11 ksi eta m x1 ux z0 i x0 lx y0 uy n j) (l_e_st_eq_landau_n_rt_cutapp2a ksi x0 lx x1 ux)))))))))))))). Time Defined. (* constant 2306 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_nis z0 x1))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_ec3e31 (l_e_st_eq_landau_n_rt_is z0 x1) (l_e_st_eq_landau_n_rt_more z0 x1) (l_e_st_eq_landau_n_rt_less z0 x1) (l_e_st_eq_landau_n_rt_satz81b z0 x1) (l_e_st_eq_landau_n_rt_rp_3140_t12 ksi eta m x1 ux z0 i x0 lx y0 uy n j)))))))))))))). Time Defined. (* constant 2307 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), l_e_st_eq_landau_n_rt_nis z0 x1))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => l_e_st_eq_landau_n_rt_rp_diffapp ksi eta z0 i (l_e_st_eq_landau_n_rt_nis z0 x1) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt eta y) => (fun (v:l_e_st_eq_landau_n_rt_more x y) => (fun (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)) => l_e_st_eq_landau_n_rt_rp_3140_t13 ksi eta m x1 ux z0 i x t y u v w))))))))))))). Time Defined. (* constant 2308 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), l_not (l_e_st_eq_landau_n_rt_in x1 (l_e_st_eq_landau_n_rt_rp_diff ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => l_imp_th3 (l_e_st_eq_landau_n_rt_in x1 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_nis x1 x1) (l_weli (l_e_st_eq_landau_n_rt_is x1 x1) (l_e_refis l_e_st_eq_landau_n_rt_rat x1)) (fun (t:l_e_st_eq_landau_n_rt_in x1 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => l_e_st_eq_landau_n_rt_rp_3140_t14 ksi eta m x1 ux x1 t)))))). Time Defined. (* constant 2309 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl z0 y0) x0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl z0 y0) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 n) y0) x0 (l_e_st_eq_landau_n_rt_ispl1 z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n) y0 j) (l_e_st_eq_landau_n_rt_satz101e x0 y0 n)))))))))))))). Time Defined. (* constant 2310 *) Definition l_e_st_eq_landau_n_rt_rp_3140_x2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_rat))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_pl u0 y0))))))))))))). Time Defined. (* constant 2311 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_rp_3140_x2 ksi eta m z0 i u0 l x0 lx y0 uy n j) x0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_pl z0 y0) x0 (l_e_st_eq_landau_n_rt_rp_3140_x2 ksi eta m z0 i u0 l x0 lx y0 uy n j) (l_e_st_eq_landau_n_rt_rp_3140_t16 ksi eta m z0 i u0 l x0 lx y0 uy n j) (l_e_st_eq_landau_n_rt_satz96c u0 z0 y0 l)))))))))))))). Time Defined. (* constant 2312 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_rp_3140_x2 ksi eta m z0 i u0 l x0 lx y0 uy n j)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_rp_satz120 ksi x0 lx (l_e_st_eq_landau_n_rt_rp_3140_x2 ksi eta m z0 i u0 l x0 lx y0 uy n j) (l_e_st_eq_landau_n_rt_rp_3140_t17 ksi eta m z0 i u0 l x0 lx y0 uy n j)))))))))))))). Time Defined. (* constant 2313 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_3140_x2 ksi eta m z0 i u0 l x0 lx y0 uy n j) y0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_ismore1 (l_e_st_eq_landau_n_rt_pl y0 u0) (l_e_st_eq_landau_n_rt_pl u0 y0) y0 (l_e_st_eq_landau_n_rt_compl y0 u0) (l_e_st_eq_landau_n_rt_satz94 y0 u0)))))))))))))). Time Defined. (* constant 2314 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_3140_x2 ksi eta m z0 i u0 l x0 lx y0 uy n j) y0 (l_e_st_eq_landau_n_rt_rp_3140_t19 ksi eta m z0 i u0 l x0 lx y0 uy n j))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_satz101g (l_e_st_eq_landau_n_rt_rp_3140_x2 ksi eta m z0 i u0 l x0 lx y0 uy n j) y0 u0 (l_e_st_eq_landau_n_rt_rp_3140_t19 ksi eta m z0 i u0 l x0 lx y0 uy n j) (l_e_st_eq_landau_n_rt_compl y0 u0)))))))))))))). Time Defined. (* constant 2315 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_rp_diff1 ksi eta u0 (l_e_st_eq_landau_n_rt_rp_3140_x2 ksi eta m z0 i u0 l x0 lx y0 uy n j) (l_e_st_eq_landau_n_rt_rp_3140_t18 ksi eta m z0 i u0 l x0 lx y0 uy n j) y0 uy (l_e_st_eq_landau_n_rt_rp_3140_t19 ksi eta m z0 i u0 l x0 lx y0 uy n j) (l_e_st_eq_landau_n_rt_rp_3140_t20 ksi eta m z0 i u0 l x0 lx y0 uy n j)))))))))))))). Time Defined. (* constant 2316 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => l_e_st_eq_landau_n_rt_rp_diffapp ksi eta z0 i (l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt eta y) => (fun (v:l_e_st_eq_landau_n_rt_more x y) => (fun (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)) => l_e_st_eq_landau_n_rt_rp_3140_t21 ksi eta m z0 i u0 l x t y u v w))))))))))))). Time Defined. (* constant 2317 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), (forall (x3:l_e_st_eq_landau_n_rt_rat), (forall (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3), (forall (l:l_e_st_eq_landau_n_rt_less x0 x3), l_e_st_eq_landau_n_rt_more x3 x0)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => (fun (x3:l_e_st_eq_landau_n_rt_rat) => (fun (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x3) => l_e_st_eq_landau_n_rt_satz83 x0 x3 l)))))))))))))). Time Defined. (* constant 2318 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t24 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), (forall (x3:l_e_st_eq_landau_n_rt_rat), (forall (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3), (forall (l:l_e_st_eq_landau_n_rt_less x0 x3), l_e_st_eq_landau_n_rt_more x3 y0)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => (fun (x3:l_e_st_eq_landau_n_rt_rat) => (fun (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x3) => l_e_st_eq_landau_n_rt_trmore x3 x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t23 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l) n)))))))))))))). Time Defined. (* constant 2319 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t25 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), (forall (x3:l_e_st_eq_landau_n_rt_rat), (forall (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3), (forall (l:l_e_st_eq_landau_n_rt_less x0 x3), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) y0) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 n) y0))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => (fun (x3:l_e_st_eq_landau_n_rt_rat) => (fun (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x3) => l_e_st_eq_landau_n_rt_ismore12 x3 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) y0) x0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 n) y0) (l_e_st_eq_landau_n_rt_satz101f x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) (l_e_st_eq_landau_n_rt_satz101f x0 y0 n) (l_e_st_eq_landau_n_rt_rp_3140_t23 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l))))))))))))))). Time Defined. (* constant 2320 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t26 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), (forall (x3:l_e_st_eq_landau_n_rt_rat), (forall (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3), (forall (l:l_e_st_eq_landau_n_rt_less x0 x3), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) (l_e_st_eq_landau_n_rt_mn x0 y0 n))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => (fun (x3:l_e_st_eq_landau_n_rt_rat) => (fun (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x3) => l_e_st_eq_landau_n_rt_satz97a (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) (l_e_st_eq_landau_n_rt_mn x0 y0 n) y0 (l_e_st_eq_landau_n_rt_rp_3140_t25 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l))))))))))))))). Time Defined. (* constant 2321 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t27 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), (forall (x3:l_e_st_eq_landau_n_rt_rat), (forall (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3), (forall (l:l_e_st_eq_landau_n_rt_less x0 x3), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) z0)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => (fun (x3:l_e_st_eq_landau_n_rt_rat) => (fun (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x3) => l_e_st_eq_landau_n_rt_ismore2 (l_e_st_eq_landau_n_rt_mn x0 y0 n) z0 (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n) j) (l_e_st_eq_landau_n_rt_rp_3140_t26 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l))))))))))))))). Time Defined. (* constant 2322 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t28 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), (forall (x3:l_e_st_eq_landau_n_rt_rat), (forall (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3), (forall (l:l_e_st_eq_landau_n_rt_less x0 x3), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) (l_e_st_eq_landau_n_rt_rp_diff ksi eta))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => (fun (x3:l_e_st_eq_landau_n_rt_rat) => (fun (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x3) => l_e_st_eq_landau_n_rt_rp_diff1 ksi eta (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) x3 lx3 y0 uy (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l))))))))))))))))). Time Defined. (* constant 2323 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t29 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), (forall (x3:l_e_st_eq_landau_n_rt_rat), (forall (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3), (forall (l:l_e_st_eq_landau_n_rt_less x0 x3), l_and (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) z0))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => (fun (x3:l_e_st_eq_landau_n_rt_rat) => (fun (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x3) => l_andi (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) z0) (l_e_st_eq_landau_n_rt_rp_3140_t28 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l) (l_e_st_eq_landau_n_rt_rp_3140_t27 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l))))))))))))))). Time Defined. (* constant 2324 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t30 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), (forall (x3:l_e_st_eq_landau_n_rt_rat), (forall (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3), (forall (l:l_e_st_eq_landau_n_rt_less x0 x3), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_more x z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => (fun (x3:l_e_st_eq_landau_n_rt_rat) => (fun (lx3:l_e_st_eq_landau_n_rt_lrt ksi x3) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x3) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_more x z0)) (l_e_st_eq_landau_n_rt_mn x3 y0 (l_e_st_eq_landau_n_rt_rp_3140_t24 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l)) (l_e_st_eq_landau_n_rt_rp_3140_t29 ksi eta m z0 i x0 lx y0 uy n j x3 lx3 l))))))))))))))). Time Defined. (* constant 2325 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t31 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (n:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_more x z0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (n:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x0 y0 n)) => l_e_st_eq_landau_n_rt_cutapp3 ksi x0 lx (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_more x z0))) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi y) => (fun (u:l_e_st_eq_landau_n_rt_less x0 y) => l_e_st_eq_landau_n_rt_rp_3140_t30 ksi eta m z0 i x0 lx y0 uy n j y t u)))))))))))))). Time Defined. (* constant 2326 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t32 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_more x z0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => l_e_st_eq_landau_n_rt_rp_diffapp ksi eta z0 i (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (l_e_st_eq_landau_n_rt_more x z0))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt eta y) => (fun (v:l_e_st_eq_landau_n_rt_more x y) => (fun (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_mn x y v)) => l_e_st_eq_landau_n_rt_rp_3140_t31 ksi eta m z0 i x t y u v w))))))))))). Time Defined. (* constant 2327 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t33 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_diff ksi eta)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => l_e_st_eq_landau_n_rt_cut2 (l_e_st_eq_landau_n_rt_rp_diff ksi eta) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t9 ksi eta m y0 ly uy x0 lx l)) (l_e_st_eq_landau_n_rt_rp_3140_t10 ksi eta m y0 ly uy x0 lx l) x1 (l_e_st_eq_landau_n_rt_rp_3140_t15 ksi eta m x1 ux1) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_less y x) => l_e_st_eq_landau_n_rt_rp_3140_t22 ksi eta m x t y u)))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) => l_e_st_eq_landau_n_rt_rp_3140_t32 ksi eta m x t))))))))))))). Time Defined. (* constant 2328 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t34 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_diff ksi eta)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_cutapp1b ksi (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => l_e_st_eq_landau_n_rt_rp_3140_t33 ksi eta m y0 ly uy x0 lx l x t))))))))))). Time Defined. (* constant 2329 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t35 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_diff ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => l_e_st_eq_landau_n_rt_cutapp3 ksi y0 ly (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_less y0 x) => l_e_st_eq_landau_n_rt_rp_3140_t34 ksi eta m y0 ly uy x t u))))))))). Time Defined. (* constant 2330 *) Definition l_e_st_eq_landau_n_rt_rp_satz140h : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_diff ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_moreapp ksi eta m (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_diff ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (v:l_e_st_eq_landau_n_rt_urt eta x) => l_e_st_eq_landau_n_rt_rp_3140_t35 ksi eta m x u v)))))). Time Defined. (* constant 2331 *) Definition l_e_st_eq_landau_n_rt_rp_3140_chi : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_cut))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_cutof (l_e_st_eq_landau_n_rt_rp_diff ksi eta) (l_e_st_eq_landau_n_rt_rp_satz140h ksi eta m)))). Time Defined. (* constant 2332 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t36 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (o:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)), l_e_st_eq_landau_n_rt_more y0 y1)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (o:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)) => l_e_st_eq_landau_n_rt_cutapp2b eta y1 ly y0 uy)))))))))))))))). Time Defined. (* constant 2333 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t37 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (o:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t36 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))) x0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (o:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t36 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) (l_e_st_eq_landau_n_rt_pl y1 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t36 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j)))) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y0) x0 (l_e_st_eq_landau_n_rt_asspl1 (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t36 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))) (l_e_st_eq_landau_n_rt_ispl2 (l_e_st_eq_landau_n_rt_pl y1 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t36 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))) y0 (l_e_st_eq_landau_n_rt_mn x0 y0 o) (l_e_st_eq_landau_n_rt_satz101c y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t36 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))) (l_e_st_eq_landau_n_rt_satz101e x0 y0 o))))))))))))))))). Time Defined. (* constant 2334 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t38 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (o:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) x0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (o:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t36 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))) x0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) (l_e_st_eq_landau_n_rt_rp_3140_t37 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j) (l_e_st_eq_landau_n_rt_satz94a (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t36 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))))))))))))))))))). Time Defined. (* constant 2335 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t39 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (o:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (o:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_pl y1 u0) (l_e_st_eq_landau_n_rt_pl u0 y1) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) i (l_e_st_eq_landau_n_rt_compl y1 u0) (l_e_st_eq_landau_n_rt_ispl1 u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1 j))))))))))))))))). Time Defined. (* constant 2336 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t40 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (o:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)), l_e_st_eq_landau_n_rt_less z0 x0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (o:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)) => l_e_st_eq_landau_n_rt_isless1 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) z0 x0 (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 o) y1) (l_e_st_eq_landau_n_rt_rp_3140_t39 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j)) (l_e_st_eq_landau_n_rt_rp_3140_t38 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))))))))))))))))). Time Defined. (* constant 2337 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t41 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (o:l_e_st_eq_landau_n_rt_more x0 y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)), l_e_st_eq_landau_n_rt_lrt ksi z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (o:l_e_st_eq_landau_n_rt_more x0 y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x0 y0 o)) => l_e_st_eq_landau_n_rt_rp_satz120 ksi x0 lx z0 (l_e_st_eq_landau_n_rt_rp_3140_t40 ksi eta m z0 lz y1 ly u0 lu i x0 lx y0 uy o j))))))))))))))))). Time Defined. (* constant 2338 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t42 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)), l_e_st_eq_landau_n_rt_lrt ksi z0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl y1 u0)) => l_e_st_eq_landau_n_rt_rp_diffapp ksi eta u0 (l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_rp_diff ksi eta) (l_e_st_eq_landau_n_rt_rp_satz140h ksi eta m) u0 lu) (l_e_st_eq_landau_n_rt_lrt ksi z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt eta y) => (fun (v:l_e_st_eq_landau_n_rt_more x y) => (fun (w:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_mn x y v)) => l_e_st_eq_landau_n_rt_rp_3140_t41 ksi eta m z0 lz y1 ly u0 lu i x t y u v w)))))))))))))))). Time Defined. (* constant 2339 *) Definition l_e_st_eq_landau_n_rt_rp_3140_a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0), l_e_st_eq_landau_n_rt_lrt ksi z0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) z0) => l_e_st_eq_landau_n_rt_rp_plapp eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) z0 lz (l_e_st_eq_landau_n_rt_lrt ksi z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_3140_t42 ksi eta m z0 lz x t y u v)))))))))). Time Defined. (* constant 2340 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t43 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_more x0 y0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_satz83 y0 x0 l))))))))). Time Defined. (* constant 2341 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t44 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), l_e_st_eq_landau_n_rt_more y2 y1))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => l_e_st_eq_landau_n_rt_cutapp2b eta y1 ly1 y2 uy2))))))))))))). Time Defined. (* constant 2342 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t45 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_more y0 y1)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_cutapp2b eta y1 ly1 y0 uy)))))))))))))). Time Defined. (* constant 2343 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t46 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_is y2 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) y1))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_tris l_e_st_eq_landau_n_rt_rat y2 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) y1) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) y1) (l_e_st_eq_landau_n_rt_satz101f y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_ispl1 (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) y1 i))))))))))))))). Time Defined. (* constant 2344 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t47 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl y2 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))) x0)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_tr4is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl y2 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) y1) (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) (l_e_st_eq_landau_n_rt_pl y1 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)))) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) y0) x0 (l_e_st_eq_landau_n_rt_ispl1 y2 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) y1) (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) (l_e_st_eq_landau_n_rt_rp_3140_t46 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) (l_e_st_eq_landau_n_rt_asspl1 (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) y1 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))) (l_e_st_eq_landau_n_rt_ispl2 (l_e_st_eq_landau_n_rt_pl y1 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))) y0 (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) (l_e_st_eq_landau_n_rt_satz101c y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))) (l_e_st_eq_landau_n_rt_satz101e x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)))))))))))))))). Time Defined. (* constant 2345 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t48 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_more x0 y2)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_ismore1 (l_e_st_eq_landau_n_rt_pl y2 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))) x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t47 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i) (l_e_st_eq_landau_n_rt_satz94 y2 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))))))))))))))))). Time Defined. (* constant 2346 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t49 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_satz101g x0 y2 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i) (l_e_st_eq_landau_n_rt_rp_3140_t47 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))))))))))))))). Time Defined. (* constant 2347 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t50 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_is y0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) y1))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_tris l_e_st_eq_landau_n_rt_rat y0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) y1) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) y1) (l_e_st_eq_landau_n_rt_satz101f y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) (l_e_st_eq_landau_n_rt_ispl1 (l_e_st_eq_landau_n_rt_mn y0 y1 (l_e_st_eq_landau_n_rt_rp_3140_t45 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) y1 (l_e_st_eq_landau_n_rt_rp_3140_t49 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)))))))))))))))). Time Defined. (* constant 2348 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t51 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_rp_ine (l_e_st_eq_landau_n_rt_rp_diff ksi eta) (l_e_st_eq_landau_n_rt_rp_satz140h ksi eta m) (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) (l_e_st_eq_landau_n_rt_rp_diff1 ksi eta (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) x0 lx y2 uy2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)))))))))))))))))). Time Defined. (* constant 2349 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t52 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt eta y1), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (uy2:l_e_st_eq_landau_n_rt_urt eta y2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt eta y1) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (uy2:l_e_st_eq_landau_n_rt_urt eta y2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y2 y1 (l_e_st_eq_landau_n_rt_rp_3140_t44 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_rp_lrtpl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) y0 y1 ly1 (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) (l_e_st_eq_landau_n_rt_rp_3140_t51 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i) (l_e_tris l_e_st_eq_landau_n_rt_rat y0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) y1) (l_e_st_eq_landau_n_rt_pl y1 (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i))) (l_e_st_eq_landau_n_rt_rp_3140_t50 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i) (l_e_st_eq_landau_n_rt_compl (l_e_st_eq_landau_n_rt_mn x0 y2 (l_e_st_eq_landau_n_rt_rp_3140_t48 ksi eta m y0 ly uy x0 lx l y1 ly1 y2 uy2 i)) y1)))))))))))))))). Time Defined. (* constant 2350 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t53 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_rp_satz132app eta (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b eta x t y u)) (l_e_st_eq_landau_n_rt_mn x0 y0 (l_e_st_eq_landau_n_rt_rp_3140_t43 ksi eta m y0 ly uy x0 lx l))) => l_e_st_eq_landau_n_rt_rp_3140_t52 ksi eta m y0 ly uy x0 lx l x t y u v)))))))))))))). Time Defined. (* constant 2351 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t54 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => l_e_st_eq_landau_n_rt_cutapp3 ksi y0 ly (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_less y0 x) => l_e_st_eq_landau_n_rt_rp_3140_t53 ksi eta m y0 ly uy x t u))))))))). Time Defined. (* constant 2352 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t55 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (ly0:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt ksi y1), (forall (uy1:l_e_st_eq_landau_n_rt_urt eta y1), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y1))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (ly0:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt ksi y1) => (fun (uy1:l_e_st_eq_landau_n_rt_urt eta y1) => l_e_st_eq_landau_n_rt_rp_3140_t54 ksi eta m y1 ly1 uy1))))))))). Time Defined. (* constant 2353 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t56 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (ly0:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (ly1:l_e_st_eq_landau_n_rt_lrt ksi y1), (forall (uy1:l_e_st_eq_landau_n_rt_urt eta y1), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (ly0:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (ly1:l_e_st_eq_landau_n_rt_lrt ksi y1) => (fun (uy1:l_e_st_eq_landau_n_rt_urt eta y1) => l_e_st_eq_landau_n_rt_rp_satz120 (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y1 (l_e_st_eq_landau_n_rt_rp_3140_t55 ksi eta m y0 ly ly0 y1 ly1 uy1) y0 (l_e_st_eq_landau_n_rt_cutapp2a eta y0 ly0 y1 uy1)))))))))). Time Defined. (* constant 2354 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t57 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), (forall (ly0:l_e_st_eq_landau_n_rt_lrt eta y0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => (fun (ly0:l_e_st_eq_landau_n_rt_lrt eta y0) => l_e_st_eq_landau_n_rt_rp_moreapp ksi eta m (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_urt eta x) => l_e_st_eq_landau_n_rt_rp_3140_t56 ksi eta m y0 ly ly0 x t u))))))))). Time Defined. (* constant 2355 *) Definition l_e_st_eq_landau_n_rt_rp_3140_b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => l_imp_th1 (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) y0) (fun (t:l_e_st_eq_landau_n_rt_lrt eta y0) => l_e_st_eq_landau_n_rt_rp_3140_t57 ksi eta m y0 ly t) (fun (t:l_e_st_eq_landau_n_rt_urt eta y0) => l_e_st_eq_landau_n_rt_rp_3140_t54 ksi eta m y0 ly t)))))). Time Defined. (* constant 2356 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t58 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) ksi (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m)) x) => l_e_st_eq_landau_n_rt_rp_3140_a ksi eta m x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => l_e_st_eq_landau_n_rt_rp_3140_b ksi eta m x t))))). Time Defined. (* constant 2357 *) Definition l_e_st_eq_landau_n_rt_rp_satz140a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_some (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta a) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_somei l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta a) ksi) (l_e_st_eq_landau_n_rt_rp_3140_chi ksi eta m) (l_e_st_eq_landau_n_rt_rp_3140_t58 ksi eta m)))). Time Defined. (* constant 2358 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t59 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_amone l_e_st_eq_landau_n_rt_cut (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta c) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (c:l_e_st_eq_landau_n_rt_cut) => (fun (d:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta c) ksi) => (fun (u:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta d) ksi) => l_e_st_eq_landau_n_rt_rp_satz140b ksi eta c d t u)))))). Time Defined. (* constant 2359 *) Definition l_e_st_eq_landau_n_rt_rp_satz140 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_one (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta a) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_onei l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta a) ksi) (l_e_st_eq_landau_n_rt_rp_3140_t59 ksi eta) (l_e_st_eq_landau_n_rt_rp_satz140a ksi eta m)))). Time Defined. (* constant 2360 *) Definition l_e_st_eq_landau_n_rt_rp_mn : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_cut))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_ind l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta a) ksi) (l_e_st_eq_landau_n_rt_rp_satz140 ksi eta m)))). Time Defined. (* constant 2361 *) Definition l_e_st_eq_landau_n_rt_rp_satz140c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_mn ksi eta m)) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_oneax l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta a) ksi) (l_e_st_eq_landau_n_rt_rp_satz140 ksi eta m)))). Time Defined. (* constant 2362 *) Definition l_e_st_eq_landau_n_rt_rp_satz140d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_mn ksi eta m))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_mn ksi eta m)) ksi (l_e_st_eq_landau_n_rt_rp_satz140c ksi eta m)))). Time Defined. (* constant 2363 *) Definition l_e_st_eq_landau_n_rt_rp_satz140e : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) eta) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) eta) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_mn ksi eta m)) ksi (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) eta) (l_e_st_eq_landau_n_rt_rp_satz140c ksi eta m)))). Time Defined. (* constant 2364 *) Definition l_e_st_eq_landau_n_rt_rp_satz140f : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) eta) ksi (l_e_st_eq_landau_n_rt_rp_satz140e ksi eta m)))). Time Defined. (* constant 2365 *) Definition l_e_st_eq_landau_n_rt_rp_satz140g : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi), l_e_st_eq_landau_n_rt_rp_is phi (l_e_st_eq_landau_n_rt_rp_mn ksi eta m)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta phi) ksi) => l_e_st_eq_landau_n_rt_rp_satz140b ksi eta phi (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) i (l_e_st_eq_landau_n_rt_rp_satz140c ksi eta m)))))). Time Defined. (* constant 2366 *) Definition l_e_st_eq_landau_n_rt_rp_3140_t60 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta), (forall (n:l_e_st_eq_landau_n_rt_rp_more eta upsilon), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl upsilon (l_e_st_eq_landau_n_rt_rp_mn ksi zeta m)) eta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more eta upsilon) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl upsilon (l_e_st_eq_landau_n_rt_rp_mn ksi zeta m)) (l_e_st_eq_landau_n_rt_rp_pl zeta (l_e_st_eq_landau_n_rt_rp_mn ksi zeta m)) ksi eta (l_e_st_eq_landau_n_rt_rp_ispl1 upsilon zeta (l_e_st_eq_landau_n_rt_rp_mn ksi zeta m) (l_e_symis l_e_st_eq_landau_n_rt_cut zeta upsilon j)) (l_e_st_eq_landau_n_rt_rp_satz140c ksi zeta m) i)))))))). Time Defined. (* constant 2367 *) Definition l_e_st_eq_landau_n_rt_rp_ismn12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta), (forall (n:l_e_st_eq_landau_n_rt_rp_more eta upsilon), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_mn ksi zeta m) (l_e_st_eq_landau_n_rt_rp_mn eta upsilon n))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more eta upsilon) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz140g eta upsilon (l_e_st_eq_landau_n_rt_rp_mn ksi zeta m) n (l_e_st_eq_landau_n_rt_rp_3140_t60 ksi eta zeta upsilon m n i j))))))))). Time Defined. (* constant 2368 *) Definition l_e_st_eq_landau_n_rt_rp_ismn1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta), (forall (n:l_e_st_eq_landau_n_rt_rp_more eta zeta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_mn ksi zeta m) (l_e_st_eq_landau_n_rt_rp_mn eta zeta n))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi zeta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more eta zeta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_ismn12 ksi eta zeta zeta m n i (l_e_refis l_e_st_eq_landau_n_rt_cut zeta))))))). Time Defined. (* constant 2369 *) Definition l_e_st_eq_landau_n_rt_rp_ismn2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more zeta ksi), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta eta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_mn zeta ksi m) (l_e_st_eq_landau_n_rt_rp_mn zeta eta n))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more zeta ksi) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta eta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_ismn12 zeta zeta ksi eta m n (l_e_refis l_e_st_eq_landau_n_rt_cut zeta) i)))))). Time Defined. (* constant 2370 *) Definition l_e_st_eq_landau_n_rt_rp_prodprop1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_and3 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0))))))). Time Defined. (* constant 2371 *) Definition l_e_st_eq_landau_n_rt_rp_prodprop : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), Prop))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x y))))). Time Defined. (* constant 2372 *) Definition l_e_st_eq_landau_n_rt_rp_prod : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_set l_e_st_eq_landau_n_rt_rat)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop ksi eta z))). Time Defined. (* constant 2373 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_and3i (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) lx ly i)))))))). Time Defined. (* constant 2374 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y) y0 (l_e_st_eq_landau_n_rt_rp_iii4_t1 ksi eta z0 x0 lx y0 ly i))))))))). Time Defined. (* constant 2375 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_rp_prodprop ksi eta z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x y)) x0 (l_e_st_eq_landau_n_rt_rp_iii4_t2 ksi eta z0 x0 lx y0 ly i))))))))). Time Defined. (* constant 2376 *) Definition l_e_st_eq_landau_n_rt_rp_prod1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop ksi eta z) z0 (l_e_st_eq_landau_n_rt_rp_iii4_t3 ksi eta z0 x0 lx y0 ly i))))))))). Time Defined. (* constant 2377 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), l_e_st_eq_landau_n_rt_rp_prodprop ksi eta z0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop ksi eta z) z0 i)))))). Time Defined. (* constant 2378 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_lrt ksi x0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0) => l_and3e1 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) py)))))))))). Time Defined. (* constant 2379 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_lrt eta y0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0) => l_and3e2 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) py)))))))))). Time Defined. (* constant 2380 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0) => l_and3e3 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_lrt eta y0) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) py)))))))))). Time Defined. (* constant 2381 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0), p)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y0) => p1 x0 (l_e_st_eq_landau_n_rt_rp_iii4_t5 ksi eta z0 i p p1 x0 px y0 py) y0 (l_e_st_eq_landau_n_rt_rp_iii4_t6 ksi eta z0 i p p1 x0 px y0 py) (l_e_st_eq_landau_n_rt_rp_iii4_t7 ksi eta z0 i p p1 x0 px y0 py))))))))))). Time Defined. (* constant 2382 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)), p)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y)) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y) px p (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x0 y) => l_e_st_eq_landau_n_rt_rp_iii4_t8 ksi eta z0 i p p1 x0 px y t)))))))))). Time Defined. (* constant 2383 *) Definition l_e_st_eq_landau_n_rt_rp_prodapp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), p)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x y)) (l_e_st_eq_landau_n_rt_rp_iii4_t4 ksi eta z0 i p p1) p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_prodprop1 ksi eta z0 x y)) => l_e_st_eq_landau_n_rt_rp_iii4_t9 ksi eta z0 i p p1 x t)))))))). Time Defined. (* constant 2384 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_less x0 x1))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_cutapp2a ksi x0 lx x1 ux))))))))))))). Time Defined. (* constant 2385 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_less y0 y1))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_cutapp2a eta y0 ly y1 uy))))))))))))). Time Defined. (* constant 2386 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_less z0 (l_e_st_eq_landau_n_rt_ts x1 y1)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_isless1 (l_e_st_eq_landau_n_rt_ts x0 y0) z0 (l_e_st_eq_landau_n_rt_ts x1 y1) (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_ts x0 y0) j) (l_e_st_eq_landau_n_rt_satz107a x0 x1 y0 y1 (l_e_st_eq_landau_n_rt_rp_4141_t1 ksi eta x1 ux y1 uy z0 i x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_rp_4141_t2 ksi eta x1 ux y1 uy z0 i x0 lx y0 ly j))))))))))))))). Time Defined. (* constant 2387 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_nis z0 (l_e_st_eq_landau_n_rt_ts x1 y1)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_ec3e31 (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x1 y1)) (l_e_st_eq_landau_n_rt_more z0 (l_e_st_eq_landau_n_rt_ts x1 y1)) (l_e_st_eq_landau_n_rt_less z0 (l_e_st_eq_landau_n_rt_ts x1 y1)) (l_e_st_eq_landau_n_rt_satz81b z0 (l_e_st_eq_landau_n_rt_ts x1 y1)) (l_e_st_eq_landau_n_rt_rp_4141_t3 ksi eta x1 ux y1 uy z0 i x0 lx y0 ly j)))))))))))))). Time Defined. (* constant 2388 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), l_e_st_eq_landau_n_rt_nis z0 (l_e_st_eq_landau_n_rt_ts x1 y1))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => l_e_st_eq_landau_n_rt_rp_prodapp ksi eta z0 i (l_e_st_eq_landau_n_rt_nis z0 (l_e_st_eq_landau_n_rt_ts x1 y1)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4141_t4 ksi eta x1 ux y1 uy z0 i x t y u v))))))))))))). Time Defined. (* constant 2389 *) Definition l_e_st_eq_landau_n_rt_rp_satz141a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), l_not (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ts x1 y1) (l_e_st_eq_landau_n_rt_rp_prod ksi eta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => l_imp_th3 (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ts x1 y1) (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_nis (l_e_st_eq_landau_n_rt_ts x1 y1) (l_e_st_eq_landau_n_rt_ts x1 y1)) (l_weli (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x1 y1) (l_e_st_eq_landau_n_rt_ts x1 y1)) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x1 y1))) (fun (t:l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ts x1 y1) (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => l_e_st_eq_landau_n_rt_rp_4141_t5 ksi eta x1 ux y1 uy (l_e_st_eq_landau_n_rt_ts x1 y1) t))))))). Time Defined. (* constant 2390 *) Definition l_e_st_eq_landau_n_rt_4141_v0 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rat)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt y0) x0)). Time Defined. (* constant 2391 *) Definition l_e_st_eq_landau_n_rt_4141_t6 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_4141_v0 x0 y0)) x0)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_4141_v0 x0 y0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt y0)) x0) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x0) x0 (l_e_st_eq_landau_n_rt_assts2 y0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt y0) x0) (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt y0)) l_e_st_eq_landau_n_rt_1rt x0 (l_e_st_eq_landau_n_rt_satz110c l_e_st_eq_landau_n_rt_1rt y0)) (l_e_st_eq_landau_n_rt_example1c x0))). Time Defined. (* constant 2392 *) Definition l_e_st_eq_landau_n_rt_satz141b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt y0) x0) (l_e_st_eq_landau_n_rt_ov x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz110g x0 y0 (l_e_st_eq_landau_n_rt_4141_v0 x0 y0) (l_e_st_eq_landau_n_rt_4141_t6 x0 y0))). Time Defined. (* constant 2393 *) Definition l_e_st_eq_landau_n_rt_satz141c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov x0 y0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt y0) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt y0) x0) (l_e_st_eq_landau_n_rt_ov x0 y0) (l_e_st_eq_landau_n_rt_satz141b x0 y0))). Time Defined. (* constant 2394 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_less z0 (l_e_st_eq_landau_n_rt_ts x0 y0)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_isless2 u0 (l_e_st_eq_landau_n_rt_ts x0 y0) z0 j l))))))))))). Time Defined. (* constant 2395 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ts x0 y0)) y0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ts x0 y0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0) y0) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt y0) y0 (l_e_st_eq_landau_n_rt_assts2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0 y0) (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0) l_e_st_eq_landau_n_rt_1rt y0 (l_e_st_eq_landau_n_rt_satz110e l_e_st_eq_landau_n_rt_1rt x0)) (l_e_st_eq_landau_n_rt_example1c y0)))))))))))). Time Defined. (* constant 2396 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ov z0 x0) y0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) z0) (l_e_st_eq_landau_n_rt_ov z0 x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ts x0 y0)) y0 (l_e_st_eq_landau_n_rt_satz141b z0 x0) (l_e_st_eq_landau_n_rt_rp_4141_t8 ksi eta u0 i z0 l x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_satz105f z0 (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_rp_4141_t7 ksi eta u0 i z0 l x0 lx y0 ly j))))))))))))). Time Defined. (* constant 2397 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_lrt eta (l_e_st_eq_landau_n_rt_ov z0 x0)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_satz120 eta y0 ly (l_e_st_eq_landau_n_rt_ov z0 x0) (l_e_st_eq_landau_n_rt_rp_4141_t9 ksi eta u0 i z0 l x0 lx y0 ly j)))))))))))). Time Defined. (* constant 2398 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_prod1 ksi eta z0 x0 lx (l_e_st_eq_landau_n_rt_ov z0 x0) (l_e_st_eq_landau_n_rt_rp_4141_t10 ksi eta u0 i z0 l x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_satz110d z0 x0)))))))))))). Time Defined. (* constant 2399 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less z0 u0), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less z0 u0) => l_e_st_eq_landau_n_rt_rp_prodapp ksi eta u0 i (l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4141_t11 ksi eta u0 i z0 l x t y u v))))))))))). Time Defined. (* constant 2400 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ts x1 y0) (l_e_st_eq_landau_n_rt_rp_prod ksi eta))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_rp_prod1 ksi eta (l_e_st_eq_landau_n_rt_ts x1 y0) x1 lx1 y0 ly (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x1 y0)))))))))))))). Time Defined. (* constant 2401 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x1 y0) (l_e_st_eq_landau_n_rt_ts x0 y0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_satz105a x1 x0 y0 (l_e_st_eq_landau_n_rt_satz83 x0 x1 l))))))))))))). Time Defined. (* constant 2402 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x1 y0) z0)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_ismore2 (l_e_st_eq_landau_n_rt_ts x0 y0) z0 (l_e_st_eq_landau_n_rt_ts x1 y0) (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_ts x0 y0) j) (l_e_st_eq_landau_n_rt_rp_4141_t14 ksi eta z0 i x0 lx y0 ly j x1 lx1 l))))))))))))). Time Defined. (* constant 2403 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_and (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ts x1 y0) (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x1 y0) z0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_andi (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ts x1 y0) (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts x1 y0) z0) (l_e_st_eq_landau_n_rt_rp_4141_t13 ksi eta z0 i x0 lx y0 ly j x1 lx1 l) (l_e_st_eq_landau_n_rt_rp_4141_t15 ksi eta z0 i x0 lx y0 ly j x1 lx1 l))))))))))))). Time Defined. (* constant 2404 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_more y z0)))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_more y z0)) (l_e_st_eq_landau_n_rt_ts x1 y0) (l_e_st_eq_landau_n_rt_rp_4141_t16 ksi eta z0 i x0 lx y0 ly j x1 lx1 l))))))))))))). Time Defined. (* constant 2405 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_more y z0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_cutapp3 ksi x0 lx (l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_more y z0))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_less x0 x) => l_e_st_eq_landau_n_rt_rp_4141_t17 ksi eta z0 i x0 lx y0 ly j x t u)))))))))))). Time Defined. (* constant 2406 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_more y z0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => l_e_st_eq_landau_n_rt_rp_prodapp ksi eta z0 i (l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_more y z0))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4141_t18 ksi eta z0 i x t y u v))))))))). Time Defined. (* constant 2407 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y1), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y1) => l_e_st_eq_landau_n_rt_cut2 (l_e_st_eq_landau_n_rt_rp_prod ksi eta) (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_rp_prod1 ksi eta (l_e_st_eq_landau_n_rt_ts x0 y0) x0 lx y0 ly (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x0 y0))) (l_e_st_eq_landau_n_rt_ts x1 y1) (l_e_st_eq_landau_n_rt_rp_satz141a ksi eta x1 ux y1 uy) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_less y x) => l_e_st_eq_landau_n_rt_rp_4141_t12 ksi eta x t y u)))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) => l_e_st_eq_landau_n_rt_rp_4141_t19 ksi eta x t)))))))))))). Time Defined. (* constant 2408 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => l_e_st_eq_landau_n_rt_cutapp1b eta (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt eta y) => l_e_st_eq_landau_n_rt_rp_4141_t20 ksi eta x0 lx y0 ly x1 ux y t)))))))))). Time Defined. (* constant 2409 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => l_e_st_eq_landau_n_rt_cutapp1b ksi (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => l_e_st_eq_landau_n_rt_rp_4141_t21 ksi eta x0 lx y0 ly x t)))))))). Time Defined. (* constant 2410 *) Definition l_e_st_eq_landau_n_rt_rp_4141_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_eq_landau_n_rt_cutapp1a eta (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta y) => l_e_st_eq_landau_n_rt_rp_4141_t22 ksi eta x0 lx y t)))))). Time Defined. (* constant 2411 *) Definition l_e_st_eq_landau_n_rt_rp_satz141 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_cutapp1a ksi (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => l_e_st_eq_landau_n_rt_rp_4141_t23 ksi eta x t)))). Time Defined. (* constant 2412 *) Definition l_e_st_eq_landau_n_rt_rp_ts : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_cutof (l_e_st_eq_landau_n_rt_rp_prod ksi eta) (l_e_st_eq_landau_n_rt_rp_satz141 ksi eta))). Time Defined. (* constant 2413 *) Definition l_e_st_eq_landau_n_rt_rp_lrtts : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_ine (l_e_st_eq_landau_n_rt_rp_prod ksi eta) (l_e_st_eq_landau_n_rt_rp_satz141 ksi eta) z0 (l_e_st_eq_landau_n_rt_rp_prod1 ksi eta z0 x0 lx y0 ly i))))))))). Time Defined. (* constant 2414 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_not (l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_isp1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_not (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_prod ksi eta))) (l_e_st_eq_landau_n_rt_ts x0 y0) z0 (l_e_st_eq_landau_n_rt_rp_satz141a ksi eta x0 ux y0 uy) i)))))))). Time Defined. (* constant 2415 *) Definition l_e_st_eq_landau_n_rt_rp_urtts : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_imp_th3 (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0) (l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta)) (l_e_st_eq_landau_n_rt_rp_iii4_t10 ksi eta z0 x0 ux y0 uy i) (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0) => l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_rp_prod ksi eta) (l_e_st_eq_landau_n_rt_rp_satz141 ksi eta) z0 t))))))))). Time Defined. (* constant 2416 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_prod ksi eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_rp_prod ksi eta) (l_e_st_eq_landau_n_rt_rp_satz141 ksi eta) z0 lz)))))). Time Defined. (* constant 2417 *) Definition l_e_st_eq_landau_n_rt_rp_tsapp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))), p)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_lrt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_lrt eta y), (forall (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)), p)))))) => l_e_st_eq_landau_n_rt_rp_prodapp ksi eta z0 (l_e_st_eq_landau_n_rt_rp_iii4_t11 ksi eta z0 lz p p1) p p1)))))). Time Defined. (* constant 2418 *) Definition l_e_st_eq_landau_n_rt_rp_ists1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ts u zeta) ksi eta i)))). Time Defined. (* constant 2419 *) Definition l_e_st_eq_landau_n_rt_rp_ists2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_cut (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ts zeta u) ksi eta i)))). Time Defined. (* constant 2420 *) Definition l_e_st_eq_landau_n_rt_rp_ists12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ists1 ksi eta zeta i) (l_e_st_eq_landau_n_rt_rp_ists2 zeta upsilon eta j))))))). Time Defined. (* constant 2421 *) Definition l_e_st_eq_landau_n_rt_rp_4142_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts y0 x0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_tris l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts y0 x0) i (l_e_st_eq_landau_n_rt_comts x0 y0)))))))))). Time Defined. (* constant 2422 *) Definition l_e_st_eq_landau_n_rt_rp_4142_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta ksi) z0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_lrtts eta ksi z0 y0 ly x0 lx (l_e_st_eq_landau_n_rt_rp_4142_t1 ksi eta z0 lz x0 lx y0 ly i)))))))))). Time Defined. (* constant 2423 *) Definition l_e_st_eq_landau_n_rt_rp_4142_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta ksi) z0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) z0) => l_e_st_eq_landau_n_rt_rp_tsapp ksi eta z0 lz (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta ksi) z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4142_t2 ksi eta z0 lz x t y u v))))))))). Time Defined. (* constant 2424 *) Definition l_e_st_eq_landau_n_rt_rp_satz142 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta ksi) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) x) => l_e_st_eq_landau_n_rt_rp_4142_t3 ksi eta x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta ksi) x) => l_e_st_eq_landau_n_rt_rp_4142_t3 eta ksi x t)))). Time Defined. (* constant 2425 *) Definition l_e_st_eq_landau_n_rt_rp_comts : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz142 ksi eta)). Time Defined. (* constant 2426 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts v0 z0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts v0 z0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_ts v0 z0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0)) i (l_e_st_eq_landau_n_rt_ists1 v0 (l_e_st_eq_landau_n_rt_ts x0 y0) z0 j) (l_e_st_eq_landau_n_rt_assts1 x0 y0 z0)))))))))))))))). Time Defined. (* constant 2427 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts v0 z0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_ts y0 z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts v0 z0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_lrtts eta zeta (l_e_st_eq_landau_n_rt_ts y0 z0) y0 ly z0 lz (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts y0 z0))))))))))))))))). Time Defined. (* constant 2428 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts v0 z0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts v0 z0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_lrtts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta) u0 x0 lx (l_e_st_eq_landau_n_rt_ts y0 z0) (l_e_st_eq_landau_n_rt_rp_4143_t2 ksi eta zeta u0 lu v0 lv z0 lz i x0 lx y0 ly j) (l_e_st_eq_landau_n_rt_rp_4143_t1 ksi eta zeta u0 lu v0 lv z0 lz i x0 lx y0 ly j)))))))))))))))). Time Defined. (* constant 2429 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts v0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts v0 z0)) => l_e_st_eq_landau_n_rt_rp_tsapp ksi eta v0 lv (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4143_t3 ksi eta zeta u0 lu v0 lv z0 lz i x t y u v))))))))))))))). Time Defined. (* constant 2430 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0) => l_e_st_eq_landau_n_rt_rp_tsapp (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta u0 lu (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt zeta y) => (fun (v:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4143_t4 ksi eta zeta u0 lu x t y u v)))))))))). Time Defined. (* constant 2431 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts y0 z0)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts y0 z0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_ts x0 v0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ts y0 z0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x0 y0) z0) i (l_e_st_eq_landau_n_rt_ists2 v0 (l_e_st_eq_landau_n_rt_ts y0 z0) x0 j) (l_e_st_eq_landau_n_rt_assts2 x0 y0 z0)))))))))))))))). Time Defined. (* constant 2432 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts y0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_ts x0 y0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts y0 z0)) => l_e_st_eq_landau_n_rt_rp_lrtts ksi eta (l_e_st_eq_landau_n_rt_ts x0 y0) x0 lx y0 ly (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x0 y0))))))))))))))))). Time Defined. (* constant 2433 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts y0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts y0 z0)) => l_e_st_eq_landau_n_rt_rp_lrtts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta u0 (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_rp_4143_t7 ksi eta zeta u0 lu x0 lx v0 lv i y0 ly z0 lz j) z0 lz (l_e_st_eq_landau_n_rt_rp_4143_t6 ksi eta zeta u0 lu x0 lx v0 lv i y0 ly z0 lz j)))))))))))))))). Time Defined. (* constant 2434 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => l_e_st_eq_landau_n_rt_rp_tsapp eta zeta v0 lv (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt zeta y) => (fun (v:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4143_t8 ksi eta zeta u0 lu x0 lx v0 lv i x t y u v))))))))))))))). Time Defined. (* constant 2435 *) Definition l_e_st_eq_landau_n_rt_rp_4143_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) u0) => l_e_st_eq_landau_n_rt_rp_tsapp ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta) u0 lu (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts eta zeta) y) => (fun (v:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4143_t9 ksi eta zeta u0 lu x t y u v)))))))))). Time Defined. (* constant 2436 *) Definition l_e_st_eq_landau_n_rt_rp_satz143 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) x) => l_e_st_eq_landau_n_rt_rp_4143_t5 ksi eta zeta x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) x) => l_e_st_eq_landau_n_rt_rp_4143_t10 ksi eta zeta x t))))). Time Defined. (* constant 2437 *) Definition l_e_st_eq_landau_n_rt_rp_assts1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz143 ksi eta zeta))). Time Defined. (* constant 2438 *) Definition l_e_st_eq_landau_n_rt_rp_assts2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_satz143 ksi eta zeta)))). Time Defined. (* constant 2439 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_ts x0 v0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x0 z0)) i (l_e_st_eq_landau_n_rt_ists2 v0 (l_e_st_eq_landau_n_rt_pl y0 z0) x0 j) (l_e_st_eq_landau_n_rt_disttp2 x0 y0 z0)))))))))))))))). Time Defined. (* constant 2440 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_ts x0 y0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_rp_lrtts ksi eta (l_e_st_eq_landau_n_rt_ts x0 y0) x0 lx y0 ly (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x0 y0))))))))))))))))). Time Defined. (* constant 2441 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_ts x0 z0)))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_rp_lrtts ksi zeta (l_e_st_eq_landau_n_rt_ts x0 z0) x0 lx z0 lz (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x0 z0))))))))))))))))). Time Defined. (* constant 2442 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl y0 z0)) => l_e_st_eq_landau_n_rt_rp_lrtpl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) u0 (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_rp_4144_t2 ksi eta zeta u0 lu x0 lx v0 lv i y0 ly z0 lz j) (l_e_st_eq_landau_n_rt_ts x0 z0) (l_e_st_eq_landau_n_rt_rp_4144_t3 ksi eta zeta u0 lu x0 lx v0 lv i y0 ly z0 lz j) (l_e_st_eq_landau_n_rt_rp_4144_t1 ksi eta zeta u0 lu x0 lx v0 lv i y0 ly z0 lz j)))))))))))))))). Time Defined. (* constant 2443 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) v0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x0 v0)) => l_e_st_eq_landau_n_rt_rp_plapp eta zeta v0 lv (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt zeta y) => (fun (v:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_4144_t4 ksi eta zeta u0 lu x0 lx v0 lv i x t y u v))))))))))))))). Time Defined. (* constant 2444 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) => l_e_st_eq_landau_n_rt_rp_tsapp ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta) u0 lu (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) y) => (fun (v:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4144_t5 ksi eta zeta u0 lu x t y u v)))))))))). Time Defined. (* constant 2445 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x1 z0)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_e_tris l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_pl v0 w0) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x1 z0)) i (l_e_st_eq_landau_n_rt_ispl12 v0 (l_e_st_eq_landau_n_rt_ts x0 y0) w0 (l_e_st_eq_landau_n_rt_ts x1 z0) j k))))))))))))))))))))). Time Defined. (* constant 2446 *) Definition l_e_st_eq_landau_n_rt_rp_4144_x2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_rat)))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_e_ite (l_e_st_eq_landau_n_rt_moreis x0 x1) l_e_st_eq_landau_n_rt_rat x0 x1)))))))))))))))))))). Time Defined. (* constant 2447 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 x1), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) x0))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 x1) => l_e_itet (l_e_st_eq_landau_n_rt_moreis x0 x1) l_e_st_eq_landau_n_rt_rat x0 x1 m))))))))))))))))))))). Time Defined. (* constant 2448 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 x1), l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 x1) => l_e_isp1 l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lrt ksi t) x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) lx (l_e_st_eq_landau_n_rt_rp_4144_t8 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k m)))))))))))))))))))))). Time Defined. (* constant 2449 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 x1), l_e_st_eq_landau_n_rt_lessis x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 x1) => l_e_st_eq_landau_n_rt_lessisi2 x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) x0 (l_e_st_eq_landau_n_rt_rp_4144_t8 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k m))))))))))))))))))))))). Time Defined. (* constant 2450 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 x1), l_e_st_eq_landau_n_rt_lessis x1 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 x1) => l_e_st_eq_landau_n_rt_satz88 x1 x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_satz84 x0 x1 m) (l_e_st_eq_landau_n_rt_rp_4144_t10 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k m)))))))))))))))))))))). Time Defined. (* constant 2451 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), (forall (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) x1))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)) => l_e_itef (l_e_st_eq_landau_n_rt_moreis x0 x1) l_e_st_eq_landau_n_rt_rat x0 x1 n))))))))))))))))))))). Time Defined. (* constant 2452 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), (forall (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)), l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)) => l_e_isp1 l_e_st_eq_landau_n_rt_rat (fun (t:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lrt ksi t) x1 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) lx1 (l_e_st_eq_landau_n_rt_rp_4144_t12 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k n)))))))))))))))))))))). Time Defined. (* constant 2453 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), (forall (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)), l_e_st_eq_landau_n_rt_lessis x1 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)) => l_e_st_eq_landau_n_rt_lessisi2 x1 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) x1 (l_e_st_eq_landau_n_rt_rp_4144_t12 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k n))))))))))))))))))))))). Time Defined. (* constant 2454 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), (forall (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)), l_e_st_eq_landau_n_rt_lessis x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)) => l_e_st_eq_landau_n_rt_lessisi1 x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_satz87b x0 x1 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_satz81j x0 x1 n) (l_e_st_eq_landau_n_rt_rp_4144_t14 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k n))))))))))))))))))))))). Time Defined. (* constant 2455 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_imp_th1 (l_e_st_eq_landau_n_rt_moreis x0 x1) (l_e_st_eq_landau_n_rt_lrt ksi (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)) (fun (t:l_e_st_eq_landau_n_rt_moreis x0 x1) => l_e_st_eq_landau_n_rt_rp_4144_t9 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k t) (fun (t:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)) => l_e_st_eq_landau_n_rt_rp_4144_t13 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k t))))))))))))))))))))). Time Defined. (* constant 2456 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lessis x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_imp_th1 (l_e_st_eq_landau_n_rt_moreis x0 x1) (l_e_st_eq_landau_n_rt_lessis x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)) (fun (t:l_e_st_eq_landau_n_rt_moreis x0 x1) => l_e_st_eq_landau_n_rt_rp_4144_t10 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k t) (fun (t:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)) => l_e_st_eq_landau_n_rt_rp_4144_t15 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k t))))))))))))))))))))). Time Defined. (* constant 2457 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lessis x1 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_imp_th1 (l_e_st_eq_landau_n_rt_moreis x0 x1) (l_e_st_eq_landau_n_rt_lessis x1 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)) (fun (t:l_e_st_eq_landau_n_rt_moreis x0 x1) => l_e_st_eq_landau_n_rt_rp_4144_t11 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k t) (fun (t:l_not (l_e_st_eq_landau_n_rt_moreis x0 x1)) => l_e_st_eq_landau_n_rt_rp_4144_t14 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k t))))))))))))))))))))). Time Defined. (* constant 2458 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_e_st_eq_landau_n_rt_rp_lrtpl eta zeta (l_e_st_eq_landau_n_rt_pl y0 z0) y0 ly z0 lz (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl y0 z0)))))))))))))))))))))). Time Defined. (* constant 2459 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_e_st_eq_landau_n_rt_rp_lrtts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_rp_4144_t16 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0) (l_e_st_eq_landau_n_rt_rp_4144_t19 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0))))))))))))))))))))))). Time Defined. (* constant 2460 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lessis (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) y0))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_e_st_eq_landau_n_rt_satz109a x0 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) y0 y0 (l_e_st_eq_landau_n_rt_rp_4144_t17 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_lessisi2 y0 y0 (l_e_refis l_e_st_eq_landau_n_rt_rat y0)))))))))))))))))))))). Time Defined. (* constant 2461 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lessis (l_e_st_eq_landau_n_rt_ts x1 z0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) z0))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_e_st_eq_landau_n_rt_satz109a x1 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) z0 z0 (l_e_st_eq_landau_n_rt_rp_4144_t18 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_lessisi2 z0 z0 (l_e_refis l_e_st_eq_landau_n_rt_rat z0)))))))))))))))))))))). Time Defined. (* constant 2462 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lessis u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0)))))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_e_st_eq_landau_n_rt_islessis12 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x1 z0)) u0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) y0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) z0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_symis l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts x1 z0)) (l_e_st_eq_landau_n_rt_rp_4144_t7 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)) (l_e_st_eq_landau_n_rt_distpt2 (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) y0 z0) (l_e_st_eq_landau_n_rt_satz100a (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) y0) (l_e_st_eq_landau_n_rt_ts x1 z0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) z0) (l_e_st_eq_landau_n_rt_rp_4144_t21 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_rp_4144_t22 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k)))))))))))))))))))))). Time Defined. (* constant 2463 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t24 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt zeta z0), (forall (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0)))))))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt zeta z0) => (fun (k:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x1 z0)) => l_orapp (l_e_st_eq_landau_n_rt_less u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0))) (l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0))) (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) (l_e_st_eq_landau_n_rt_rp_4144_t23 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (fun (t:l_e_st_eq_landau_n_rt_less u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0))) => l_e_st_eq_landau_n_rt_rp_satz120 (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0)) (l_e_st_eq_landau_n_rt_rp_4144_t20 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) u0 t) (fun (t:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0))) => l_e_isp1 l_e_st_eq_landau_n_rt_rat (fun (u:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_4144_x2 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) (l_e_st_eq_landau_n_rt_pl y0 z0)) u0 (l_e_st_eq_landau_n_rt_rp_4144_t20 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x1 lx1 z0 lz k) t))))))))))))))))))))). Time Defined. (* constant 2464 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t25 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt eta y0), (forall (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0))))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt eta y0) => (fun (j:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_tsapp ksi zeta w0 lw (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt zeta y) => (fun (v:l_e_st_eq_landau_n_rt_is w0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4144_t24 ksi eta zeta u0 lu v0 lv w0 lw i x0 lx y0 ly j x t y u v)))))))))))))))))))). Time Defined. (* constant 2465 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t26 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0), (forall (w0:l_e_st_eq_landau_n_rt_rat), (forall (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0), (forall (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) v0) => (fun (w0:l_e_st_eq_landau_n_rt_rat) => (fun (lw:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) w0) => (fun (i:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl v0 w0)) => l_e_st_eq_landau_n_rt_rp_tsapp ksi eta v0 lv (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta y) => (fun (v:l_e_st_eq_landau_n_rt_is v0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4144_t25 ksi eta zeta u0 lu v0 lv w0 lw i x t y u v))))))))))))))). Time Defined. (* constant 2466 *) Definition l_e_st_eq_landau_n_rt_rp_4144_t27 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) u0) => l_e_st_eq_landau_n_rt_rp_plapp (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) u0 lu (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi eta) x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) y) => (fun (v:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_4144_t26 ksi eta zeta u0 lu x t y u v)))))))))). Time Defined. (* constant 2467 *) Definition l_e_st_eq_landau_n_rt_rp_satz144 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) x) => l_e_st_eq_landau_n_rt_rp_4144_t6 ksi eta zeta x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) x) => l_e_st_eq_landau_n_rt_rp_4144_t27 ksi eta zeta x t))))). Time Defined. (* constant 2468 *) Definition l_e_st_eq_landau_n_rt_rp_disttp1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta (l_e_st_eq_landau_n_rt_rp_pl ksi eta)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_satz144 zeta ksi eta) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta eta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_comts zeta ksi) (l_e_st_eq_landau_n_rt_rp_comts zeta eta))))). Time Defined. (* constant 2469 *) Definition l_e_st_eq_landau_n_rt_rp_disttp2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz144 ksi eta zeta))). Time Defined. (* constant 2470 *) Definition l_e_st_eq_landau_n_rt_rp_distpt1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi eta) zeta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_disttp1 ksi eta zeta)))). Time Defined. (* constant 2471 *) Definition l_e_st_eq_landau_n_rt_rp_distpt2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_pl eta zeta)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta)) (l_e_st_eq_landau_n_rt_rp_disttp2 ksi eta zeta)))). Time Defined. (* constant 2472 *) Definition l_e_st_eq_landau_n_rt_rp_4145_phi : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_mn ksi eta m)))). Time Defined. (* constant 2473 *) Definition l_e_st_eq_landau_n_rt_rp_4145_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz140d ksi eta m)))). Time Defined. (* constant 2474 *) Definition l_e_st_eq_landau_n_rt_rp_4145_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m) zeta)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m)) zeta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m) zeta)) (l_e_st_eq_landau_n_rt_rp_ists1 ksi (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m)) zeta (l_e_st_eq_landau_n_rt_rp_4145_t1 ksi eta zeta m)) (l_e_st_eq_landau_n_rt_rp_disttp1 eta (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m) zeta))))). Time Defined. (* constant 2475 *) Definition l_e_st_eq_landau_n_rt_rp_satz145a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_ismore1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m) zeta)) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m) zeta)) (l_e_st_eq_landau_n_rt_rp_4145_t2 ksi eta zeta m)) (l_e_st_eq_landau_n_rt_rp_satz133 (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_4145_phi ksi eta zeta m) zeta)))))). Time Defined. (* constant 2476 *) Definition l_e_st_eq_landau_n_rt_rp_satz145b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_ists1 ksi eta zeta i)))). Time Defined. (* constant 2477 *) Definition l_e_st_eq_landau_n_rt_rp_satz145c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz145a eta ksi zeta (l_e_st_eq_landau_n_rt_rp_satz122 ksi eta l)))))). Time Defined. (* constant 2478 *) Definition l_e_st_eq_landau_n_rt_rp_satz145d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta eta) (l_e_st_eq_landau_n_rt_rp_comts ksi zeta) (l_e_st_eq_landau_n_rt_rp_comts eta zeta) (l_e_st_eq_landau_n_rt_rp_satz145a ksi eta zeta m))))). Time Defined. (* constant 2479 *) Definition l_e_st_eq_landau_n_rt_rp_satz145e : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_ists2 ksi eta zeta i)))). Time Defined. (* constant 2480 *) Definition l_e_st_eq_landau_n_rt_rp_satz145f : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta eta) (l_e_st_eq_landau_n_rt_rp_comts ksi zeta) (l_e_st_eq_landau_n_rt_rp_comts eta zeta) (l_e_st_eq_landau_n_rt_rp_satz145c ksi eta zeta l))))). Time Defined. (* constant 2481 *) Definition l_e_st_eq_landau_n_rt_rp_satz145g : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_ismore2 (l_e_st_eq_landau_n_rt_rp_ts ksi upsilon) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ists1 ksi eta upsilon i) (l_e_st_eq_landau_n_rt_rp_satz145d zeta upsilon ksi m))))))). Time Defined. (* constant 2482 *) Definition l_e_st_eq_landau_n_rt_rp_satz145h : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts upsilon eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ts upsilon eta) (l_e_st_eq_landau_n_rt_rp_comts ksi zeta) (l_e_st_eq_landau_n_rt_rp_comts eta upsilon) (l_e_st_eq_landau_n_rt_rp_satz145g ksi eta zeta upsilon i m))))))). Time Defined. (* constant 2483 *) Definition l_e_st_eq_landau_n_rt_rp_satz145j : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less zeta upsilon) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_ts ksi upsilon) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ists1 ksi eta upsilon i) (l_e_st_eq_landau_n_rt_rp_satz145f zeta upsilon ksi l))))))). Time Defined. (* constant 2484 *) Definition l_e_st_eq_landau_n_rt_rp_satz145k : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (l:l_e_st_eq_landau_n_rt_rp_less zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts upsilon eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (l:l_e_st_eq_landau_n_rt_rp_less zeta upsilon) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ts upsilon eta) (l_e_st_eq_landau_n_rt_rp_comts ksi zeta) (l_e_st_eq_landau_n_rt_rp_comts eta upsilon) (l_e_st_eq_landau_n_rt_rp_satz145j ksi eta zeta upsilon i l))))))). Time Defined. (* constant 2485 *) Definition l_e_st_eq_landau_n_rt_rp_4146_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_or3 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz123a ksi eta))). Time Defined. (* constant 2486 *) Definition l_e_st_eq_landau_n_rt_rp_4146_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_ec3 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)))). Time Defined. (* constant 2487 *) Definition l_e_st_eq_landau_n_rt_rp_satz146a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)), l_e_st_eq_landau_n_rt_rp_more ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) => l_ec3_th11 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_4146_t1 ksi eta zeta) (l_e_st_eq_landau_n_rt_rp_4146_t2 ksi eta zeta) (fun (u:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145b ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145a ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145c ksi eta zeta u) m)))). Time Defined. (* constant 2488 *) Definition l_e_st_eq_landau_n_rt_rp_satz146b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)), l_e_st_eq_landau_n_rt_rp_is ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) => l_ec3_th10 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_4146_t1 ksi eta zeta) (l_e_st_eq_landau_n_rt_rp_4146_t2 ksi eta zeta) (fun (u:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145b ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145a ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145c ksi eta zeta u) i)))). Time Defined. (* constant 2489 *) Definition l_e_st_eq_landau_n_rt_rp_satz146c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)), l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) => l_ec3_th12 (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta)) (l_e_st_eq_landau_n_rt_rp_4146_t1 ksi eta zeta) (l_e_st_eq_landau_n_rt_rp_4146_t2 ksi eta zeta) (fun (u:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145b ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145a ksi eta zeta u) (fun (u:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145c ksi eta zeta u) l)))). Time Defined. (* constant 2490 *) Definition l_e_st_eq_landau_n_rt_rp_satz146d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta)), l_e_st_eq_landau_n_rt_rp_more ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta)) => l_e_st_eq_landau_n_rt_rp_satz146a ksi eta zeta (l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta eta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_comts zeta ksi) (l_e_st_eq_landau_n_rt_rp_comts zeta eta) m))))). Time Defined. (* constant 2491 *) Definition l_e_st_eq_landau_n_rt_rp_satz146e : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta)), l_e_st_eq_landau_n_rt_rp_is ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta)) => l_e_st_eq_landau_n_rt_rp_satz146b ksi eta zeta (l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_comts ksi zeta) i (l_e_st_eq_landau_n_rt_rp_comts zeta eta)))))). Time Defined. (* constant 2492 *) Definition l_e_st_eq_landau_n_rt_rp_satz146f : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta)), l_e_st_eq_landau_n_rt_rp_less ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts zeta eta)) => l_e_st_eq_landau_n_rt_rp_satz146c ksi eta zeta (l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_ts zeta ksi) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts zeta eta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_comts zeta ksi) (l_e_st_eq_landau_n_rt_rp_comts zeta eta) l))))). Time Defined. (* constant 2493 *) Definition l_e_st_eq_landau_n_rt_rp_4147_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz145a ksi eta zeta m)))))). Time Defined. (* constant 2494 *) Definition l_e_st_eq_landau_n_rt_rp_4147_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_ts zeta eta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts upsilon eta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_comts zeta eta) (l_e_st_eq_landau_n_rt_rp_comts upsilon eta) (l_e_st_eq_landau_n_rt_rp_satz145a zeta upsilon eta n))))))). Time Defined. (* constant 2495 *) Definition l_e_st_eq_landau_n_rt_rp_satz147 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_trmore (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_4147_t1 ksi eta zeta upsilon m n) (l_e_st_eq_landau_n_rt_rp_4147_t2 ksi eta zeta upsilon m n))))))). Time Defined. (* constant 2496 *) Definition l_e_st_eq_landau_n_rt_rp_satz147a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz147 eta ksi upsilon zeta (l_e_st_eq_landau_n_rt_rp_satz122 ksi eta l) (l_e_st_eq_landau_n_rt_rp_satz122 zeta upsilon k)))))))). Time Defined. (* constant 2497 *) Definition l_e_st_eq_landau_n_rt_rp_satz148a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_orapp (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon)) m (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz147 ksi eta zeta upsilon t n) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_satz145g ksi eta zeta upsilon t n))))))). Time Defined. (* constant 2498 *) Definition l_e_st_eq_landau_n_rt_rp_satz148b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => l_orapp (l_e_st_eq_landau_n_rt_rp_more zeta upsilon) (l_e_st_eq_landau_n_rt_rp_is zeta upsilon) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon)) n (fun (t:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz147 ksi eta zeta upsilon m t) (fun (t:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz145h zeta upsilon ksi eta t m))))))). Time Defined. (* constant 2499 *) Definition l_e_st_eq_landau_n_rt_rp_satz148c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_less zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_less zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz148a eta ksi upsilon zeta (l_e_st_eq_landau_n_rt_rp_satz125 ksi eta l) (l_e_st_eq_landau_n_rt_rp_satz122 zeta upsilon k)))))))). Time Defined. (* constant 2500 *) Definition l_e_st_eq_landau_n_rt_rp_satz148d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis zeta upsilon), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz148b eta ksi upsilon zeta (l_e_st_eq_landau_n_rt_rp_satz122 ksi eta l) (l_e_st_eq_landau_n_rt_rp_satz125 zeta upsilon k)))))))). Time Defined. (* constant 2501 *) Definition l_e_st_eq_landau_n_rt_rp_4149_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_st_eq_landau_n_rt_rp_moreisi2 (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ists12 ksi eta zeta upsilon i j))))))))). Time Defined. (* constant 2502 *) Definition l_e_st_eq_landau_n_rt_rp_4149_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), (forall (o:l_e_st_eq_landau_n_rt_rp_more zeta upsilon), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => (fun (o:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_moreisi1 (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_satz148a ksi eta zeta upsilon m o))))))))). Time Defined. (* constant 2503 *) Definition l_e_st_eq_landau_n_rt_rp_4149_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_orapp (l_e_st_eq_landau_n_rt_rp_more zeta upsilon) (l_e_st_eq_landau_n_rt_rp_is zeta upsilon) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon)) n (fun (t:l_e_st_eq_landau_n_rt_rp_more zeta upsilon) => l_e_st_eq_landau_n_rt_rp_4149_t2 ksi eta zeta upsilon m n i t) (fun (t:l_e_st_eq_landau_n_rt_rp_is zeta upsilon) => l_e_st_eq_landau_n_rt_rp_4149_t1 ksi eta zeta upsilon m n i t)))))))). Time Defined. (* constant 2504 *) Definition l_e_st_eq_landau_n_rt_rp_4149_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), (forall (o:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => (fun (o:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_moreisi1 (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_satz148b ksi eta zeta upsilon o n)))))))). Time Defined. (* constant 2505 *) Definition l_e_st_eq_landau_n_rt_rp_satz149 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta), (forall (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis ksi eta) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreis zeta upsilon) => l_orapp (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_is ksi eta) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon)) m (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_4149_t4 ksi eta zeta upsilon m n t) (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_st_eq_landau_n_rt_rp_4149_t3 ksi eta zeta upsilon m n t))))))). Time Defined. (* constant 2506 *) Definition l_e_st_eq_landau_n_rt_rp_satz149a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (upsilon:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta), (forall (k:l_e_st_eq_landau_n_rt_rp_lessis zeta upsilon), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_ts eta upsilon))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (upsilon:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessis ksi eta) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessis zeta upsilon) => l_e_st_eq_landau_n_rt_rp_satz124 (l_e_st_eq_landau_n_rt_rp_ts eta upsilon) (l_e_st_eq_landau_n_rt_rp_ts ksi zeta) (l_e_st_eq_landau_n_rt_rp_satz149 eta ksi upsilon zeta (l_e_st_eq_landau_n_rt_rp_satz125 ksi eta l) (l_e_st_eq_landau_n_rt_rp_satz125 zeta upsilon k)))))))). Time Defined. (* constant 2507 *) Definition l_e_st_eq_landau_n_rt_ratset : (forall (r0:l_e_st_eq_landau_n_rt_rat), l_e_st_set l_e_st_eq_landau_n_rt_rat). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less x r0)). Time Defined. (* constant 2508 *) Definition l_e_st_eq_landau_n_rt_4150_t1 : (forall (r0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less x r0)). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz90 r0). Time Defined. (* constant 2509 *) Definition l_e_st_eq_landau_n_rt_4150_t2 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 r0), l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 r0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less x r0) x0 l))). Time Defined. (* constant 2510 *) Definition l_e_st_eq_landau_n_rt_4150_t3 : (forall (r0:l_e_st_eq_landau_n_rt_rat), l_not (l_e_st_eq_landau_n_rt_less r0 r0)). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => l_ec3e13 (l_e_st_eq_landau_n_rt_is r0 r0) (l_e_st_eq_landau_n_rt_more r0 r0) (l_e_st_eq_landau_n_rt_less r0 r0) (l_e_st_eq_landau_n_rt_satz81b r0 r0) (l_e_refis l_e_st_eq_landau_n_rt_rat r0)). Time Defined. (* constant 2511 *) Definition l_e_st_eq_landau_n_rt_4150_t4 : (forall (r0:l_e_st_eq_landau_n_rt_rat), l_not (l_e_st_eq_landau_n_rt_in r0 (l_e_st_eq_landau_n_rt_ratset r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => l_imp_th3 (l_e_st_eq_landau_n_rt_in r0 (l_e_st_eq_landau_n_rt_ratset r0)) (l_e_st_eq_landau_n_rt_less r0 r0) (l_e_st_eq_landau_n_rt_4150_t3 r0) (fun (t:l_e_st_eq_landau_n_rt_in r0 (l_e_st_eq_landau_n_rt_ratset r0)) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less x r0) r0 t)). Time Defined. (* constant 2512 *) Definition l_e_st_eq_landau_n_rt_4150_t5 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)), l_e_st_eq_landau_n_rt_less x0 r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less x r0) x0 i))). Time Defined. (* constant 2513 *) Definition l_e_st_eq_landau_n_rt_4150_t6 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (k:l_e_st_eq_landau_n_rt_less x1 x0), l_e_st_eq_landau_n_rt_in x1 (l_e_st_eq_landau_n_rt_ratset r0)))))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (k:l_e_st_eq_landau_n_rt_less x1 x0) => l_e_st_eq_landau_n_rt_4150_t2 r0 x1 (l_e_st_eq_landau_n_rt_trless x1 x0 r0 k (l_e_st_eq_landau_n_rt_4150_t5 r0 x0 i))))))). Time Defined. (* constant 2514 *) Definition l_e_st_eq_landau_n_rt_4150_t7 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x0 x) (l_e_st_eq_landau_n_rt_less x r0))))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)) => l_e_st_eq_landau_n_rt_satz91 x0 r0 (l_e_st_eq_landau_n_rt_4150_t5 r0 x0 i)))). Time Defined. (* constant 2515 *) Definition l_e_st_eq_landau_n_rt_4150_t8 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x0 x1) (l_e_st_eq_landau_n_rt_less x1 r0)), l_e_st_eq_landau_n_rt_less x0 x1))))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x0 x1) (l_e_st_eq_landau_n_rt_less x1 r0)) => l_ande1 (l_e_st_eq_landau_n_rt_less x0 x1) (l_e_st_eq_landau_n_rt_less x1 r0) a))))). Time Defined. (* constant 2516 *) Definition l_e_st_eq_landau_n_rt_4150_t9 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x0 x1) (l_e_st_eq_landau_n_rt_less x1 r0)), l_e_st_eq_landau_n_rt_less x1 r0))))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x0 x1) (l_e_st_eq_landau_n_rt_less x1 r0)) => l_ande2 (l_e_st_eq_landau_n_rt_less x0 x1) (l_e_st_eq_landau_n_rt_less x1 r0) a))))). Time Defined. (* constant 2517 *) Definition l_e_st_eq_landau_n_rt_4150_t10 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x0 x1) (l_e_st_eq_landau_n_rt_less x1 r0)), l_and (l_e_st_eq_landau_n_rt_in x1 (l_e_st_eq_landau_n_rt_ratset r0)) (l_e_st_eq_landau_n_rt_more x1 x0)))))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x0 x1) (l_e_st_eq_landau_n_rt_less x1 r0)) => l_andi (l_e_st_eq_landau_n_rt_in x1 (l_e_st_eq_landau_n_rt_ratset r0)) (l_e_st_eq_landau_n_rt_more x1 x0) (l_e_st_eq_landau_n_rt_4150_t2 r0 x1 (l_e_st_eq_landau_n_rt_4150_t9 r0 x0 i x1 a)) (l_e_st_eq_landau_n_rt_satz83 x0 x1 (l_e_st_eq_landau_n_rt_4150_t8 r0 x0 i x1 a))))))). Time Defined. (* constant 2518 *) Definition l_e_st_eq_landau_n_rt_4150_t11 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_ratset r0)) (l_e_st_eq_landau_n_rt_more x x0))))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_ratset r0)) => l_some_th6 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x0 x) (l_e_st_eq_landau_n_rt_less x r0)) (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_ratset r0)) (l_e_st_eq_landau_n_rt_more x x0)) (l_e_st_eq_landau_n_rt_4150_t7 r0 x0 i) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_and (l_e_st_eq_landau_n_rt_less x0 x) (l_e_st_eq_landau_n_rt_less x r0)) => l_e_st_eq_landau_n_rt_4150_t10 r0 x0 i x t))))). Time Defined. (* constant 2519 *) Definition l_e_st_eq_landau_n_rt_4150_t12 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 r0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_ratset r0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 r0) => l_e_st_eq_landau_n_rt_cut2 (l_e_st_eq_landau_n_rt_ratset r0) x0 (l_e_st_eq_landau_n_rt_4150_t2 r0 x0 l) r0 (l_e_st_eq_landau_n_rt_4150_t4 r0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_ratset r0)) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_less y x) => l_e_st_eq_landau_n_rt_4150_t6 r0 x t y u)))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_ratset r0)) => l_e_st_eq_landau_n_rt_4150_t11 r0 x t))))). Time Defined. (* constant 2520 *) Definition l_e_st_eq_landau_n_rt_satz150 : (forall (r0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_ratset r0)). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_less x r0) (l_e_st_eq_landau_n_rt_4150_t1 r0) (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_ratset r0)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_less x r0) => l_e_st_eq_landau_n_rt_4150_t12 r0 x t))). Time Defined. (* constant 2521 *) Definition l_e_st_eq_landau_n_rt_rp_rpofrt : (forall (r0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_cutof (l_e_st_eq_landau_n_rt_ratset r0) (l_e_st_eq_landau_n_rt_satz150 r0)). Time Defined. (* constant 2522 *) Definition l_e_st_eq_landau_n_rt_rp_lrtrpofrt : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 r0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt r0) x0))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 r0) => l_e_st_eq_landau_n_rt_rp_ine (l_e_st_eq_landau_n_rt_ratset r0) (l_e_st_eq_landau_n_rt_satz150 r0) x0 (l_e_st_eq_landau_n_rt_4150_t2 r0 x0 l)))). Time Defined. (* constant 2523 *) Definition l_e_st_eq_landau_n_rt_rp_lrtrpofrte : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt r0) x0), l_e_st_eq_landau_n_rt_less x0 r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt r0) x0) => l_e_st_eq_landau_n_rt_4150_t5 r0 x0 (l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_ratset r0) (l_e_st_eq_landau_n_rt_satz150 r0) x0 lx)))). Time Defined. (* constant 2524 *) Definition l_e_st_eq_landau_n_rt_rp_iii4_t12 : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 r0), l_not (l_e_st_eq_landau_n_rt_less x0 r0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 r0) => l_e_st_eq_landau_n_rt_satz81c x0 r0 m))). Time Defined. (* constant 2525 *) Definition l_e_st_eq_landau_n_rt_rp_urtrpofrt : (forall (r0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 r0), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_rpofrt r0) x0))). exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 r0) => l_imp_th3 (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt r0) x0) (l_e_st_eq_landau_n_rt_less x0 r0) (l_e_st_eq_landau_n_rt_rp_iii4_t12 r0 x0 m) (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt r0) x0) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte r0 x0 t)))). Time Defined. (* constant 2526 *) Definition l_e_st_eq_landau_n_rt_rp_1rp : l_e_st_eq_landau_n_rt_cut. exact (l_e_st_eq_landau_n_rt_rp_rpofrt l_e_st_eq_landau_n_rt_1rt). Time Defined. (* constant 2527 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_less y0 l_e_st_eq_landau_n_rt_1rt)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte l_e_st_eq_landau_n_rt_1rt y0 ly)))))))). Time Defined. (* constant 2528 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_less z0 x0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_ts x0 y0) z0 (l_e_st_eq_landau_n_rt_ts x0 l_e_st_eq_landau_n_rt_1rt) x0 (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_ts x0 y0) i) (l_e_st_eq_landau_n_rt_example1a x0) (l_e_st_eq_landau_n_rt_satz105f y0 l_e_st_eq_landau_n_rt_1rt x0 (l_e_st_eq_landau_n_rt_rp_4151_t1 ksi z0 lz x0 lx y0 ly i)))))))))). Time Defined. (* constant 2529 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp y0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)), l_e_st_eq_landau_n_rt_lrt ksi z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp y0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 y0)) => l_e_st_eq_landau_n_rt_rp_satz120 ksi x0 lx z0 (l_e_st_eq_landau_n_rt_rp_4151_t2 ksi z0 lz x0 lx y0 ly i))))))))). Time Defined. (* constant 2530 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) z0), l_e_st_eq_landau_n_rt_lrt ksi z0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) z0) => l_e_st_eq_landau_n_rt_rp_tsapp ksi l_e_st_eq_landau_n_rt_rp_1rp z0 lz (l_e_st_eq_landau_n_rt_lrt ksi z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4151_t3 ksi z0 lz x t y u v)))))))). Time Defined. (* constant 2531 *) Definition l_e_st_eq_landau_n_rt_rp_4151_y1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_rat)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) x0)))))). Time Defined. (* constant 2532 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_rp_4151_y1 ksi x0 lx x1 lx1 l) l_e_st_eq_landau_n_rt_1rt)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) x1) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_4151_y1 ksi x0 lx x1 lx1 l) (l_e_st_eq_landau_n_rt_satz110e l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_satz105f x0 x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) l))))))). Time Defined. (* constant 2533 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_4151_y1 ksi x0 lx x1 lx1 l))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_4151_y1 ksi x0 lx x1 lx1 l) (l_e_st_eq_landau_n_rt_rp_4151_t5 ksi x0 lx x1 lx1 l))))))). Time Defined. (* constant 2534 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x1 (l_e_st_eq_landau_n_rt_rp_4151_y1 ksi x0 lx x1 lx1 l)) x0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x1 (l_e_st_eq_landau_n_rt_rp_4151_y1 ksi x0 lx x1 lx1 l)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) x0) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x0) x0 (l_e_st_eq_landau_n_rt_assts2 x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) x0) (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_ts x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) l_e_st_eq_landau_n_rt_1rt x0 (l_e_st_eq_landau_n_rt_satz110c l_e_st_eq_landau_n_rt_1rt x1)) (l_e_st_eq_landau_n_rt_example1c x0))))))). Time Defined. (* constant 2535 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x0 x1), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) x0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x0 x1) => l_e_st_eq_landau_n_rt_rp_lrtts ksi l_e_st_eq_landau_n_rt_rp_1rp x0 x1 lx1 (l_e_st_eq_landau_n_rt_rp_4151_y1 ksi x0 lx x1 lx1 l) (l_e_st_eq_landau_n_rt_rp_4151_t6 ksi x0 lx x1 lx1 l) (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x1 (l_e_st_eq_landau_n_rt_rp_4151_y1 ksi x0 lx x1 lx1 l)) x0 (l_e_st_eq_landau_n_rt_rp_4151_t7 ksi x0 lx x1 lx1 l)))))))). Time Defined. (* constant 2536 *) Definition l_e_st_eq_landau_n_rt_rp_4151_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) x0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_eq_landau_n_rt_cutapp3 ksi x0 lx (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) x0) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi y) => (fun (u:l_e_st_eq_landau_n_rt_less x0 y) => l_e_st_eq_landau_n_rt_rp_4151_t8 ksi x0 lx y t u)))))). Time Defined. (* constant 2537 *) Definition l_e_st_eq_landau_n_rt_rp_satz151 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) ksi). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) ksi (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) x) => l_e_st_eq_landau_n_rt_rp_4151_t4 ksi x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => l_e_st_eq_landau_n_rt_rp_4151_t9 ksi x t))). Time Defined. (* constant 2538 *) Definition l_e_st_eq_landau_n_rt_rp_satz151a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) ksi (l_e_st_eq_landau_n_rt_rp_satz151 ksi)). Time Defined. (* constant 2539 *) Definition l_e_st_eq_landau_n_rt_rp_satz151b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_1rp ksi) ksi). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_1rp ksi) (l_e_st_eq_landau_n_rt_rp_ts ksi l_e_st_eq_landau_n_rt_rp_1rp) ksi (l_e_st_eq_landau_n_rt_rp_comts l_e_st_eq_landau_n_rt_rp_1rp ksi) (l_e_st_eq_landau_n_rt_rp_satz151 ksi)). Time Defined. (* constant 2540 *) Definition l_e_st_eq_landau_n_rt_rp_satz151c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_1rp ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_1rp ksi) ksi (l_e_st_eq_landau_n_rt_rp_satz151b ksi)). Time Defined. (* constant 2541 *) Definition l_e_st_eq_landau_n_rt_rp_4152_invprop1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), Prop))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_less y0 x0)))). Time Defined. (* constant 2542 *) Definition l_e_st_eq_landau_n_rt_rp_4152_invprop2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_and3 (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x)) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0))))). Time Defined. (* constant 2543 *) Definition l_e_st_eq_landau_n_rt_rp_4152_invprop : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), Prop)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x))). Time Defined. (* constant 2544 *) Definition l_e_st_eq_landau_n_rt_rp_4152_inv : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_set l_e_st_eq_landau_n_rt_rat). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop ksi z)). Time Defined. (* constant 2545 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 y0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_andi (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_less y0 x0) uy l)))))))). Time Defined. (* constant 2546 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x) y0 (l_e_st_eq_landau_n_rt_rp_4152_t1 ksi z0 x0 ux y0 uy l i))))))))). Time Defined. (* constant 2547 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_and3i (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x)) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) ux (l_e_st_eq_landau_n_rt_rp_4152_t2 ksi z0 x0 ux y0 uy l i) i)))))))). Time Defined. (* constant 2548 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_rp_4152_invprop ksi z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x) x0 (l_e_st_eq_landau_n_rt_rp_4152_t3 ksi z0 x0 ux y0 uy l i))))))))). Time Defined. (* constant 2549 *) Definition l_e_st_eq_landau_n_rt_rp_4152_inv1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (z:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop ksi z) z0 (l_e_st_eq_landau_n_rt_rp_4152_t4 ksi z0 x0 ux y0 uy l i))))))))). Time Defined. (* constant 2550 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), l_e_st_eq_landau_n_rt_rp_4152_invprop ksi z0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop ksi x) z0 i))))). Time Defined. (* constant 2551 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0), l_e_st_eq_landau_n_rt_urt ksi x0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0) => l_and3e1 (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x)) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) px))))))). Time Defined. (* constant 2552 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0) => l_and3e2 (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x)) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) px))))))). Time Defined. (* constant 2553 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0) => l_and3e3 (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x)) (l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) px))))))). Time Defined. (* constant 2554 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 y0), l_e_st_eq_landau_n_rt_urt ksi y0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 y0) => l_ande1 (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_less y0 x0) py))))))))). Time Defined. (* constant 2555 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 y0), l_e_st_eq_landau_n_rt_less y0 x0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 y0) => l_ande2 (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_less y0 x0) py))))))))). Time Defined. (* constant 2556 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 y0), p))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 y0) => p1 x0 (l_e_st_eq_landau_n_rt_rp_4152_t6 ksi z0 i p p1 x0 px) y0 (l_e_st_eq_landau_n_rt_rp_4152_t9 ksi z0 i p p1 x0 px y0 py) (l_e_st_eq_landau_n_rt_rp_4152_t10 ksi z0 i p p1 x0 px y0 py) (l_e_st_eq_landau_n_rt_rp_4152_t8 ksi z0 i p p1 x0 px)))))))))). Time Defined. (* constant 2557 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0), p))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x0) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x) (l_e_st_eq_landau_n_rt_rp_4152_t7 ksi z0 i p p1 x0 px) p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_4152_invprop1 ksi x0 x) => l_e_st_eq_landau_n_rt_rp_4152_t11 ksi z0 i p p1 x0 px x t))))))))). Time Defined. (* constant 2558 *) Definition l_e_st_eq_landau_n_rt_rp_4152_invapp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))), p))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_urt ksi x), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_urt ksi y), (forall (v:l_e_st_eq_landau_n_rt_less y x), (forall (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)), p))))))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x) (l_e_st_eq_landau_n_rt_rp_4152_t5 ksi z0 i p p1) p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_4152_invprop2 ksi z0 x) => l_e_st_eq_landau_n_rt_rp_4152_t12 ksi z0 i p p1 x t))))))). Time Defined. (* constant 2559 *) Definition l_e_st_eq_landau_n_rt_rp_4152_2x0 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_rat))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_e_st_eq_landau_n_rt_pl x0 x0))). Time Defined. (* constant 2560 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_less x0 (l_e_st_eq_landau_n_rt_rp_4152_2x0 ksi x0 ux)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_e_st_eq_landau_n_rt_satz94a x0 x0))). Time Defined. (* constant 2561 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_rp_4152_2x0 ksi x0 ux)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_e_st_eq_landau_n_rt_rp_satz119a ksi x0 ux (l_e_st_eq_landau_n_rt_rp_4152_2x0 ksi x0 ux) (l_e_st_eq_landau_n_rt_rp_4152_t13 ksi x0 ux)))). Time Defined. (* constant 2562 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_4152_2x0 ksi x0 ux)) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_e_st_eq_landau_n_rt_rp_4152_inv1 ksi (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_4152_2x0 ksi x0 ux)) (l_e_st_eq_landau_n_rt_rp_4152_2x0 ksi x0 ux) (l_e_st_eq_landau_n_rt_rp_4152_t14 ksi x0 ux) x0 ux (l_e_st_eq_landau_n_rt_rp_4152_t13 ksi x0 ux) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_4152_2x0 ksi x0 ux)))))). Time Defined. (* constant 2563 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_nis x0 x1))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_imp_th3 (l_e_st_eq_landau_n_rt_is x0 x1) (l_e_st_eq_landau_n_rt_lrt ksi x0) ux (fun (t:l_e_st_eq_landau_n_rt_is x0 x1) => l_e_isp1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lrt ksi x) x1 x0 lx t)))))). Time Defined. (* constant 2564 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0) l_e_st_eq_landau_n_rt_1rt))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_e_st_eq_landau_n_rt_satz110e l_e_st_eq_landau_n_rt_1rt x0))))). Time Defined. (* constant 2565 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) x1) l_e_st_eq_landau_n_rt_1rt))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_e_st_eq_landau_n_rt_satz110e l_e_st_eq_landau_n_rt_1rt x1))))). Time Defined. (* constant 2566 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x1) l_e_st_eq_landau_n_rt_1rt)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x1) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) x1) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) x1 i) (l_e_st_eq_landau_n_rt_rp_4152_t18 ksi x1 lx x0 ux))))))). Time Defined. (* constant 2567 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)), l_e_st_eq_landau_n_rt_is x0 x1)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) => l_e_st_eq_landau_n_rt_satz110b l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0 x1 (l_e_st_eq_landau_n_rt_rp_4152_t17 ksi x1 lx x0 ux) (l_e_st_eq_landau_n_rt_rp_4152_t19 ksi x1 lx x0 ux i))))))). Time Defined. (* constant 2568 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_nis (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_imp_th3 (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) (l_e_st_eq_landau_n_rt_is x0 x1) (l_e_st_eq_landau_n_rt_rp_4152_t16 ksi x1 lx x0 ux) (fun (t:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) => l_e_st_eq_landau_n_rt_rp_4152_t20 ksi x1 lx x0 ux t)))))). Time Defined. (* constant 2569 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (i:l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), l_con)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (i:l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => l_e_st_eq_landau_n_rt_rp_4152_invapp ksi (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) i l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt ksi y) => (fun (v:l_e_st_eq_landau_n_rt_less y x) => (fun (w:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)) => l_e_st_eq_landau_n_rt_rp_4152_t21 ksi x1 lx x t (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x) w))))))))))). Time Defined. (* constant 2570 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x1), l_not (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (t:l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => l_e_st_eq_landau_n_rt_rp_4152_t22 ksi x1 lx t)))). Time Defined. (* constant 2571 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t24 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_less u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_isless2 z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) u0 j l)))))))). Time Defined. (* constant 2572 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t25 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0) (l_e_st_eq_landau_n_rt_ts u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ts u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0)) (l_e_st_eq_landau_n_rt_satz110e l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_satz110d l_e_st_eq_landau_n_rt_1rt u0))))))))). Time Defined. (* constant 2573 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t26 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts u0 x0) (l_e_st_eq_landau_n_rt_ts u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0) (l_e_st_eq_landau_n_rt_ts u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0)) (l_e_st_eq_landau_n_rt_ts u0 x0) (l_e_st_eq_landau_n_rt_rp_4152_t25 ksi z0 i u0 l x0 ux j) (l_e_st_eq_landau_n_rt_satz105c u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x0 (l_e_st_eq_landau_n_rt_rp_4152_t24 ksi z0 i u0 l x0 ux j)))))))))). Time Defined. (* constant 2574 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t27 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 u0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) u0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_ts u0 x0) (l_e_st_eq_landau_n_rt_ts x0 u0) (l_e_st_eq_landau_n_rt_ts u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) u0) (l_e_st_eq_landau_n_rt_comts u0 x0) (l_e_st_eq_landau_n_rt_comts u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0)) (l_e_st_eq_landau_n_rt_rp_4152_t26 ksi z0 i u0 l x0 ux j))))))))). Time Defined. (* constant 2575 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t28 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_less x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_satz106c x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) u0 (l_e_st_eq_landau_n_rt_rp_4152_t27 ksi z0 i u0 l x0 ux j))))))))). Time Defined. (* constant 2576 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t29 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_rp_satz119a ksi x0 ux (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) (l_e_st_eq_landau_n_rt_rp_4152_t28 ksi z0 i u0 l x0 ux j))))))))). Time Defined. (* constant 2577 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t30 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) u0) l_e_st_eq_landau_n_rt_1rt)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_satz110e l_e_st_eq_landau_n_rt_1rt u0)))))))). Time Defined. (* constant 2578 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t31 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_satz110g l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) u0 (l_e_st_eq_landau_n_rt_rp_4152_t30 ksi z0 i u0 l x0 ux j))))))))). Time Defined. (* constant 2579 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t32 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_rp_4152_inv1 ksi u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) (l_e_st_eq_landau_n_rt_rp_4152_t29 ksi z0 i u0 l x0 ux j) x0 ux (l_e_st_eq_landau_n_rt_rp_4152_t28 ksi z0 i u0 l x0 ux j) (l_e_st_eq_landau_n_rt_rp_4152_t31 ksi z0 i u0 l x0 ux j))))))))). Time Defined. (* constant 2580 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t33 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less u0 z0), l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less u0 z0) => l_e_st_eq_landau_n_rt_rp_4152_invapp ksi z0 i (l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt ksi y) => (fun (v:l_e_st_eq_landau_n_rt_less y x) => (fun (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)) => l_e_st_eq_landau_n_rt_rp_4152_t32 ksi z0 i u0 l x t w))))))))))). Time Defined. (* constant 2581 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t34 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x1 x) (l_e_st_eq_landau_n_rt_less x x0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_e_st_eq_landau_n_rt_satz91 x1 x0 l))))))))). Time Defined. (* constant 2582 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t35 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_less x1 x2))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_ande1 (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0) a))))))))))). Time Defined. (* constant 2583 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t36 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_urt ksi x2))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_e_st_eq_landau_n_rt_rp_satz119a ksi x1 ux1 x2 (l_e_st_eq_landau_n_rt_rp_4152_t35 ksi z0 i x0 ux x1 ux1 l j x2 a)))))))))))). Time Defined. (* constant 2584 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t37 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_e_st_eq_landau_n_rt_rp_4152_inv1 ksi (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) x2 (l_e_st_eq_landau_n_rt_rp_4152_t36 ksi z0 i x0 ux x1 ux1 l j x2 a) x1 ux1 (l_e_st_eq_landau_n_rt_rp_4152_t35 ksi z0 i x0 ux x1 ux1 l j x2 a) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2))))))))))))). Time Defined. (* constant 2585 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t38 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_less x2 x0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_ande2 (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0) a))))))))))). Time Defined. (* constant 2586 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t39 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) (l_e_st_eq_landau_n_rt_ts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_e_tris l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2)) (l_e_st_eq_landau_n_rt_satz110c l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_satz110d l_e_st_eq_landau_n_rt_1rt x2)))))))))))). Time Defined. (* constant 2587 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t40 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) (l_e_st_eq_landau_n_rt_ts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) (l_e_st_eq_landau_n_rt_ts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2)) (l_e_st_eq_landau_n_rt_ts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) (l_e_st_eq_landau_n_rt_rp_4152_t39 ksi z0 i x0 ux x1 ux1 l j x2 a) (l_e_st_eq_landau_n_rt_satz105c x2 x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_rp_4152_t38 ksi z0 i x0 ux x1 ux1 l j x2 a))))))))))))). Time Defined. (* constant 2588 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t41 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x2) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) x2)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_ts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) x2) (l_e_st_eq_landau_n_rt_ts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) x2) (l_e_st_eq_landau_n_rt_comts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) (l_e_st_eq_landau_n_rt_comts x2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2)) (l_e_st_eq_landau_n_rt_rp_4152_t40 ksi z0 i x0 ux x1 ux1 l j x2 a)))))))))))). Time Defined. (* constant 2589 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t42 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_e_st_eq_landau_n_rt_satz106c (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) x2 (l_e_st_eq_landau_n_rt_rp_4152_t41 ksi z0 i x0 ux x1 ux1 l j x2 a)))))))))))). Time Defined. (* constant 2590 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t43 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) z0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_e_st_eq_landau_n_rt_ismore2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) j) (l_e_st_eq_landau_n_rt_satz83 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) (l_e_st_eq_landau_n_rt_rp_4152_t42 ksi z0 i x0 ux x1 ux1 l j x2 a))))))))))))). Time Defined. (* constant 2591 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t44 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_and (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) z0)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_andi (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) z0) (l_e_st_eq_landau_n_rt_rp_4152_t37 ksi z0 i x0 ux x1 ux1 l j x2 a) (l_e_st_eq_landau_n_rt_rp_4152_t43 ksi z0 i x0 ux x1 ux1 l j x2 a)))))))))))). Time Defined. (* constant 2592 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t45 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (l_e_st_eq_landau_n_rt_more x z0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less x1 x2) (l_e_st_eq_landau_n_rt_less x2 x0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (l_e_st_eq_landau_n_rt_more x z0)) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x2) (l_e_st_eq_landau_n_rt_rp_4152_t44 ksi z0 i x0 ux x1 ux1 l j x2 a)))))))))))). Time Defined. (* constant 2593 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t46 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux1:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (l:l_e_st_eq_landau_n_rt_less x1 x0), (forall (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (l_e_st_eq_landau_n_rt_more x z0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux1:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (l:l_e_st_eq_landau_n_rt_less x1 x0) => (fun (j:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0)) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less x1 x) (l_e_st_eq_landau_n_rt_less x x0)) (l_e_st_eq_landau_n_rt_rp_4152_t34 ksi z0 i x0 ux x1 ux1 l j) (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (l_e_st_eq_landau_n_rt_more x z0))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_and (l_e_st_eq_landau_n_rt_less x1 x) (l_e_st_eq_landau_n_rt_less x x0)) => l_e_st_eq_landau_n_rt_rp_4152_t45 ksi z0 i x0 ux x1 ux1 l j x t))))))))))). Time Defined. (* constant 2594 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t47 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (l_e_st_eq_landau_n_rt_more x z0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in z0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => l_e_st_eq_landau_n_rt_rp_4152_invapp ksi z0 i (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (l_e_st_eq_landau_n_rt_more x z0))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt ksi y) => (fun (v:l_e_st_eq_landau_n_rt_less y x) => (fun (w:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)) => l_e_st_eq_landau_n_rt_rp_4152_t46 ksi z0 i x t y u v w))))))))). Time Defined. (* constant 2595 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t48 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_e_st_eq_landau_n_rt_cut2 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_pl y0 y0)) (l_e_st_eq_landau_n_rt_rp_4152_t15 ksi y0 uy) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x0) (l_e_st_eq_landau_n_rt_rp_4152_t23 ksi x0 lx) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_less y x) => l_e_st_eq_landau_n_rt_rp_4152_t33 ksi x t y u)))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) => l_e_st_eq_landau_n_rt_rp_4152_t47 ksi x t))))))). Time Defined. (* constant 2596 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t49 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_eq_landau_n_rt_cutapp1b ksi (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => l_e_st_eq_landau_n_rt_rp_4152_t48 ksi x0 lx x t))))). Time Defined. (* constant 2597 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t50 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_cutapp1a ksi (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_4152_inv ksi)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => l_e_st_eq_landau_n_rt_rp_4152_t49 ksi x t))). Time Defined. (* constant 2598 *) Definition l_e_st_eq_landau_n_rt_rp_4152_chi : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_cutof (l_e_st_eq_landau_n_rt_rp_4152_inv ksi) (l_e_st_eq_landau_n_rt_rp_4152_t50 ksi)). Time Defined. (* constant 2599 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t51 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)), l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) => l_e_tris l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_ts x0 u0) (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) i (l_e_st_eq_landau_n_rt_ists2 u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) x0 j)))))))))))). Time Defined. (* constant 2600 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t52 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)), l_e_st_eq_landau_n_rt_less x0 x1))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) => l_e_st_eq_landau_n_rt_cutapp2a ksi x0 lx x1 ux))))))))))). Time Defined. (* constant 2601 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t53 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)), l_e_st_eq_landau_n_rt_less z0 l_e_st_eq_landau_n_rt_1rt))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) z0 (l_e_st_eq_landau_n_rt_ts x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) l_e_st_eq_landau_n_rt_1rt (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) (l_e_st_eq_landau_n_rt_rp_4152_t51 ksi z0 lz x0 lx u0 lu i x1 ux j)) (l_e_st_eq_landau_n_rt_satz110c l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_satz105c x0 x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1) (l_e_st_eq_landau_n_rt_rp_4152_t52 ksi z0 lz x0 lx u0 lu i x1 ux j))))))))))))). Time Defined. (* constant 2602 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t54 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x1), (forall (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)), l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp z0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x1) => (fun (j:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x1)) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt l_e_st_eq_landau_n_rt_1rt z0 (l_e_st_eq_landau_n_rt_rp_4152_t53 ksi z0 lz x0 lx u0 lu i x1 ux j)))))))))))). Time Defined. (* constant 2603 *) Definition l_e_st_eq_landau_n_rt_rp_4152_r1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)), l_e_st_eq_landau_n_rt_in u0 (l_e_st_eq_landau_n_rt_rp_4152_inv ksi))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)) => l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_rp_4152_inv ksi) (l_e_st_eq_landau_n_rt_rp_4152_t50 ksi) u0 lu)))))))). Time Defined. (* constant 2604 *) Definition l_e_st_eq_landau_n_rt_rp_4152_r2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)), l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp z0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x0 u0)) => l_e_st_eq_landau_n_rt_rp_4152_invapp ksi u0 (l_e_st_eq_landau_n_rt_rp_4152_r1 ksi z0 lz x0 lx u0 lu i) (l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt ksi y) => (fun (v:l_e_st_eq_landau_n_rt_less y x) => (fun (w:l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt x)) => l_e_st_eq_landau_n_rt_rp_4152_t54 ksi z0 lz x0 lx u0 lu i x t w)))))))))))))). Time Defined. (* constant 2605 *) Definition l_e_st_eq_landau_n_rt_rp_4152_r3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0), l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp z0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) z0) => l_e_st_eq_landau_n_rt_rp_tsapp ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) z0 lz (l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_4152_r2 ksi z0 lz x t y u v)))))))). Time Defined. (* constant 2606 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t55 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), l_e_st_eq_landau_n_rt_less u0 l_e_st_eq_landau_n_rt_1rt))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte l_e_st_eq_landau_n_rt_1rt u0 lu))). Time Defined. (* constant 2607 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t56 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), l_e_st_eq_landau_n_rt_more l_e_st_eq_landau_n_rt_1rt u0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => l_e_st_eq_landau_n_rt_satz83 u0 l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_rp_4152_t55 ksi u0 lu)))). Time Defined. (* constant 2608 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t57 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), l_e_st_eq_landau_n_rt_more x2 x1))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => l_e_st_eq_landau_n_rt_cutapp2b ksi x1 lx1 x2 ux2))))))))). Time Defined. (* constant 2609 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t58 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_less x0 x2)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_cutapp2a ksi x0 lx x2 ux2)))))))))). Time Defined. (* constant 2610 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t59 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x2))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_satz105f x0 x2 (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) (l_e_st_eq_landau_n_rt_rp_4152_t58 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2611 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t60 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x2))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_isless1 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0) (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x2) (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0) i) (l_e_st_eq_landau_n_rt_rp_4152_t59 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2612 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t61 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x2) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) x1))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_tr4is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x2) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) u0) x2) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x2) x2 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) x1) (l_e_st_eq_landau_n_rt_distpt1 (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) u0 x2) (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) u0) l_e_st_eq_landau_n_rt_1rt x2 (l_e_st_eq_landau_n_rt_satz101e l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu))) (l_e_st_eq_landau_n_rt_example1c x2) (l_e_st_eq_landau_n_rt_satz101f x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)))))))))))). Time Defined. (* constant 2613 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t62 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x2) (l_e_st_eq_landau_n_rt_ts u0 x2)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_satz96c (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x2) (l_e_st_eq_landau_n_rt_ts u0 x2) (l_e_st_eq_landau_n_rt_rp_4152_t60 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2614 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t63 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) x1))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_isless2 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x2) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) x1) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_rp_4152_t61 ksi u0 lu x0 lx x1 lx1 x2 ux2 i) (l_e_st_eq_landau_n_rt_rp_4152_t62 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2615 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t64 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts u0 x2) (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2))) (l_e_st_eq_landau_n_rt_pl x1 (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts u0 x2) (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2))) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) x1) (l_e_st_eq_landau_n_rt_pl x1 (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2))) (l_e_st_eq_landau_n_rt_compl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_compl (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) x1) (l_e_st_eq_landau_n_rt_rp_4152_t63 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2616 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t65 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts u0 x2) x1)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_satz97c (l_e_st_eq_landau_n_rt_ts u0 x2) x1 (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_rp_4152_t64 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2617 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t66 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) (l_e_st_eq_landau_n_rt_ts u0 x2)) x2)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) (l_e_st_eq_landau_n_rt_ts u0 x2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) u0) x2) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x2) x2 (l_e_st_eq_landau_n_rt_assts2 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) u0 x2) (l_e_st_eq_landau_n_rt_ists1 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) u0) l_e_st_eq_landau_n_rt_1rt x2 (l_e_st_eq_landau_n_rt_satz110e l_e_st_eq_landau_n_rt_1rt u0)) (l_e_st_eq_landau_n_rt_example1c x2))))))))))). Time Defined. (* constant 2618 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t67 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_less x2 (l_e_st_eq_landau_n_rt_ov x1 u0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) (l_e_st_eq_landau_n_rt_ts u0 x2)) x2 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) x1) (l_e_st_eq_landau_n_rt_ov x1 u0) (l_e_st_eq_landau_n_rt_rp_4152_t66 ksi u0 lu x0 lx x1 lx1 x2 ux2 i) (l_e_st_eq_landau_n_rt_satz141b x1 u0) (l_e_st_eq_landau_n_rt_satz105f (l_e_st_eq_landau_n_rt_ts u0 x2) x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt u0) (l_e_st_eq_landau_n_rt_rp_4152_t65 ksi u0 lu x0 lx x1 lx1 x2 ux2 i)))))))))))). Time Defined. (* constant 2619 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t68 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_ov x1 u0))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_rp_satz119a ksi x2 ux2 (l_e_st_eq_landau_n_rt_ov x1 u0) (l_e_st_eq_landau_n_rt_rp_4152_t67 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2620 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t69 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov x1 u0) u0) x1)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_satz110e x1 u0)))))))))). Time Defined. (* constant 2621 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t70 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_tr3is l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_ov x1 (l_e_st_eq_landau_n_rt_ov x1 u0)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0)) x1) (l_e_st_eq_landau_n_rt_ts x1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0))) (l_e_st_eq_landau_n_rt_satz110g x1 (l_e_st_eq_landau_n_rt_ov x1 u0) u0 (l_e_st_eq_landau_n_rt_rp_4152_t69 ksi u0 lu x0 lx x1 lx1 x2 ux2 i)) (l_e_st_eq_landau_n_rt_satz141c x1 (l_e_st_eq_landau_n_rt_ov x1 u0)) (l_e_st_eq_landau_n_rt_comts (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0)) x1))))))))))). Time Defined. (* constant 2622 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t71 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0)) (l_e_st_eq_landau_n_rt_rp_4152_inv ksi))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_rp_4152_inv1 ksi (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0)) (l_e_st_eq_landau_n_rt_ov x1 u0) (l_e_st_eq_landau_n_rt_rp_4152_t68 ksi u0 lu x0 lx x1 lx1 x2 ux2 i) x2 ux2 (l_e_st_eq_landau_n_rt_rp_4152_t67 ksi u0 lu x0 lx x1 lx1 x2 ux2 i) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0))))))))))))). Time Defined. (* constant 2623 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t72 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_rp_ine (l_e_st_eq_landau_n_rt_rp_4152_inv ksi) (l_e_st_eq_landau_n_rt_rp_4152_t50 ksi) (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0)) (l_e_st_eq_landau_n_rt_rp_4152_t71 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2624 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t73 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (ux2:l_e_st_eq_landau_n_rt_urt ksi x2), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) u0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt ksi x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (ux2:l_e_st_eq_landau_n_rt_urt ksi x2) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn x2 x1 (l_e_st_eq_landau_n_rt_rp_4152_t57 ksi u0 lu x0 lx x1 lx1 x2 ux2)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_rp_lrtts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) u0 x1 lx1 (l_e_st_eq_landau_n_rt_ov l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_ov x1 u0)) (l_e_st_eq_landau_n_rt_rp_4152_t72 ksi u0 lu x0 lx x1 lx1 x2 ux2 i) (l_e_st_eq_landau_n_rt_rp_4152_t70 ksi u0 lu x0 lx x1 lx1 x2 ux2 i))))))))))). Time Defined. (* constant 2625 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t74 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) u0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_eq_landau_n_rt_rp_satz132app ksi (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) u0) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_urt ksi y) => (fun (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn y x (l_e_st_eq_landau_n_rt_cutapp2b ksi x t y u)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_mn l_e_st_eq_landau_n_rt_1rt u0 (l_e_st_eq_landau_n_rt_rp_4152_t56 ksi u0 lu)) x0)) => l_e_st_eq_landau_n_rt_rp_4152_t73 ksi u0 lu x0 lx x t y u v)))))))))). Time Defined. (* constant 2626 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t75 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) u0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp u0) => l_e_st_eq_landau_n_rt_cutapp1a ksi (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => l_e_st_eq_landau_n_rt_rp_4152_t74 ksi u0 lu x t))))). Time Defined. (* constant 2627 *) Definition l_e_st_eq_landau_n_rt_rp_4152_t76 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) l_e_st_eq_landau_n_rt_rp_1rp). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) l_e_st_eq_landau_n_rt_rp_1rp (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_4152_chi ksi)) x) => l_e_st_eq_landau_n_rt_rp_4152_r3 ksi x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt l_e_st_eq_landau_n_rt_rp_1rp x) => l_e_st_eq_landau_n_rt_rp_4152_t75 ksi x t))). Time Defined. (* constant 2628 *) Definition l_e_st_eq_landau_n_rt_rp_satz152 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_some (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi c) l_e_st_eq_landau_n_rt_rp_1rp)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_somei l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi t) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_4152_chi ksi) (l_e_st_eq_landau_n_rt_rp_4152_t76 ksi)). Time Defined. (* constant 2629 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more phi psi), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more phi psi) => l_e_st_eq_landau_n_rt_rp_satz145d phi psi eta m))))). Time Defined. (* constant 2630 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more phi psi), l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more phi psi) => l_ec3e21 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_4153_t1 ksi eta phi psi m)))))). Time Defined. (* constant 2631 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less phi psi), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less phi psi) => l_e_st_eq_landau_n_rt_rp_satz145f phi psi eta l))))). Time Defined. (* constant 2632 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less phi psi), l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less phi psi) => l_ec3e31 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_4153_t3 ksi eta phi psi l)))))). Time Defined. (* constant 2633 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_nis phi psi), l_or (l_e_st_eq_landau_n_rt_rp_more phi psi) (l_e_st_eq_landau_n_rt_rp_less phi psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_nis phi psi) => l_or3_th1 (l_e_st_eq_landau_n_rt_rp_is phi psi) (l_e_st_eq_landau_n_rt_rp_more phi psi) (l_e_st_eq_landau_n_rt_rp_less phi psi) (l_e_st_eq_landau_n_rt_rp_satz123a phi psi) n))))). Time Defined. (* constant 2634 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_nis phi psi), l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_nis phi psi) => l_orapp (l_e_st_eq_landau_n_rt_rp_more phi psi) (l_e_st_eq_landau_n_rt_rp_less phi psi) (l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_st_eq_landau_n_rt_rp_4153_t5 ksi eta phi psi n) (fun (t:l_e_st_eq_landau_n_rt_rp_more phi psi) => l_e_st_eq_landau_n_rt_rp_4153_t2 ksi eta phi psi t) (fun (t:l_e_st_eq_landau_n_rt_rp_less phi psi) => l_e_st_eq_landau_n_rt_rp_4153_t4 ksi eta phi psi t)))))). Time Defined. (* constant 2635 *) Definition l_e_st_eq_landau_n_rt_rp_satz153b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (psi:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta phi) ksi), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta psi) ksi), l_e_st_eq_landau_n_rt_rp_is phi psi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (psi:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta phi) ksi) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta psi) ksi) => l_imp_th7 (l_e_st_eq_landau_n_rt_rp_is phi psi) (l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_weli (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi)) (l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts eta phi) (l_e_st_eq_landau_n_rt_rp_ts eta psi) ksi i j)) (fun (t:l_e_st_eq_landau_n_rt_rp_nis phi psi) => l_e_st_eq_landau_n_rt_rp_4153_t6 ksi eta phi psi t))))))). Time Defined. (* constant 2636 *) Definition l_e_st_eq_landau_n_rt_rp_4153_chi : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (tau:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta tau) l_e_st_eq_landau_n_rt_rp_1rp), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (tau:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta tau) l_e_st_eq_landau_n_rt_rp_1rp) => l_e_st_eq_landau_n_rt_rp_ts tau ksi)))). Time Defined. (* constant 2637 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (tau:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta tau) l_e_st_eq_landau_n_rt_rp_1rp), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_4153_chi ksi eta tau i)) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (tau:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta tau) l_e_st_eq_landau_n_rt_rp_1rp) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_4153_chi ksi eta tau i)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts eta tau) ksi) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_1rp ksi) ksi (l_e_st_eq_landau_n_rt_rp_assts2 eta tau ksi) (l_e_st_eq_landau_n_rt_rp_ists1 (l_e_st_eq_landau_n_rt_rp_ts eta tau) l_e_st_eq_landau_n_rt_rp_1rp ksi i) (l_e_st_eq_landau_n_rt_rp_satz151b ksi))))). Time Defined. (* constant 2638 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (tau:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta tau) l_e_st_eq_landau_n_rt_rp_1rp), l_e_st_eq_landau_n_rt_rp_some (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) ksi))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (tau:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta tau) l_e_st_eq_landau_n_rt_rp_1rp) => l_somei l_e_st_eq_landau_n_rt_cut (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) ksi) (l_e_st_eq_landau_n_rt_rp_4153_chi ksi eta tau i) (l_e_st_eq_landau_n_rt_rp_4153_t7 ksi eta tau i))))). Time Defined. (* constant 2639 *) Definition l_e_st_eq_landau_n_rt_rp_satz153a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_some (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_someapp l_e_st_eq_landau_n_rt_cut (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_satz152 eta) (l_e_st_eq_landau_n_rt_rp_some (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) ksi)) (fun (c:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) l_e_st_eq_landau_n_rt_rp_1rp) => l_e_st_eq_landau_n_rt_rp_4153_t8 ksi eta c t)))). Time Defined. (* constant 2640 *) Definition l_e_st_eq_landau_n_rt_rp_4153_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_amone l_e_st_eq_landau_n_rt_cut (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (c:l_e_st_eq_landau_n_rt_cut) => (fun (d:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) ksi) => (fun (u:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta d) ksi) => l_e_st_eq_landau_n_rt_rp_satz153b ksi eta c d t u)))))). Time Defined. (* constant 2641 *) Definition l_e_st_eq_landau_n_rt_rp_satz153 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_one (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_onei l_e_st_eq_landau_n_rt_cut (fun (c:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta c) ksi) (l_e_st_eq_landau_n_rt_rp_4153_t9 ksi eta) (l_e_st_eq_landau_n_rt_rp_satz153a ksi eta))). Time Defined. (* constant 2642 *) Definition l_e_st_eq_landau_n_rt_rp_ov : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_ind l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta a) ksi) (l_e_st_eq_landau_n_rt_rp_satz153 ksi eta))). Time Defined. (* constant 2643 *) Definition l_e_st_eq_landau_n_rt_rp_satz153c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_ov ksi eta)) ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_oneax l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta a) ksi) (l_e_st_eq_landau_n_rt_rp_satz153 ksi eta))). Time Defined. (* constant 2644 *) Definition l_e_st_eq_landau_n_rt_rp_satz153d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_ov ksi eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_ov ksi eta)) ksi (l_e_st_eq_landau_n_rt_rp_satz153c ksi eta))). Time Defined. (* constant 2645 *) Definition l_e_st_eq_landau_n_rt_rp_satz153e : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ov ksi eta) eta) ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ov ksi eta) eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_ov ksi eta)) ksi (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_ov ksi eta) eta) (l_e_st_eq_landau_n_rt_rp_satz153c ksi eta))). Time Defined. (* constant 2646 *) Definition l_e_st_eq_landau_n_rt_rp_satz153f : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ov ksi eta) eta))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ov ksi eta) eta) ksi (l_e_st_eq_landau_n_rt_rp_satz153e ksi eta))). Time Defined. (* constant 2647 *) Definition l_e_st_eq_landau_n_rt_rp_satz153g : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (phi:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta phi) ksi), l_e_st_eq_landau_n_rt_rp_is phi (l_e_st_eq_landau_n_rt_rp_ov ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (phi:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts eta phi) ksi) => l_e_st_eq_landau_n_rt_rp_satz153b ksi eta phi (l_e_st_eq_landau_n_rt_rp_ov ksi eta) i (l_e_st_eq_landau_n_rt_rp_satz153c ksi eta))))). Time Defined. (* constant 2648 *) Definition l_e_st_eq_landau_n_rt_rp_ratrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), Prop). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_image l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi). Time Defined. (* constant 2649 *) Definition l_e_st_eq_landau_n_rt_rp_ratrpi : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_rpofrt x0)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_imagei l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) x0). Time Defined. (* constant 2650 *) Definition l_e_st_eq_landau_n_rt_rp_rpofnt : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_cut). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rtofn x)). Time Defined. (* constant 2651 *) Definition l_e_st_eq_landau_n_rt_rp_natrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), Prop). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt x) ksi). Time Defined. (* constant 2652 *) Definition l_e_st_eq_landau_n_rt_rp_natrpi : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofnt x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_imagei l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt y) x). Time Defined. (* constant 2653 *) Definition l_e_st_eq_landau_n_rt_rp_iii5_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (x:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofnt x)), l_e_st_eq_landau_n_rt_rp_ratrp ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofnt x)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofrt y)) (l_e_st_eq_landau_n_rt_rtofn x) i)))). Time Defined. (* constant 2654 *) Definition l_e_st_eq_landau_n_rt_rp_lemmaiii5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp ksi), l_e_st_eq_landau_n_rt_rp_ratrp ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp ksi) => l_someapp l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofnt x)) n (l_e_st_eq_landau_n_rt_rp_ratrp ksi) (fun (x:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofnt x)) => l_e_st_eq_landau_n_rt_rp_iii5_t1 ksi n x t)))). Time Defined. (* constant 2655 *) Definition l_e_st_eq_landau_n_rt_rp_5154_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt x0 y0 (l_e_st_eq_landau_n_rt_satz82 x0 y0 m)))). Time Defined. (* constant 2656 *) Definition l_e_st_eq_landau_n_rt_rp_5154_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_rp_urtrpofrt y0 y0 (l_e_st_eq_landau_n_rt_moreisi2 y0 y0 (l_e_refis l_e_st_eq_landau_n_rt_rat y0))))). Time Defined. (* constant 2657 *) Definition l_e_st_eq_landau_n_rt_rp_5154_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_and (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) y0) (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_andi (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) y0) (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) y0) (l_e_st_eq_landau_n_rt_rp_5154_t1 x0 y0 m) (l_e_st_eq_landau_n_rt_rp_5154_t2 x0 y0 m)))). Time Defined. (* constant 2658 *) Definition l_e_st_eq_landau_n_rt_rp_satz154a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) x) (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) x)) y0 (l_e_st_eq_landau_n_rt_rp_5154_t3 x0 y0 m)))). Time Defined. (* constant 2659 *) Definition l_e_st_eq_landau_n_rt_rp_satz154b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_isf l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) x0 y0 i))). Time Defined. (* constant 2660 *) Definition l_e_st_eq_landau_n_rt_rp_satz154c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_satz154a y0 x0 (l_e_st_eq_landau_n_rt_satz83 x0 y0 l))))). Time Defined. (* constant 2661 *) Definition l_e_st_eq_landau_n_rt_rp_5154_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_or3 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz81a x0 y0)). Time Defined. (* constant 2662 *) Definition l_e_st_eq_landau_n_rt_rp_5154_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_ec3 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0))). Time Defined. (* constant 2663 *) Definition l_e_st_eq_landau_n_rt_rp_satz154d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_more x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_ec3_th11 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_5154_t4 x0 y0) (l_e_st_eq_landau_n_rt_rp_5154_t5 x0 y0) (fun (u:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154b x0 y0 u) (fun (u:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154a x0 y0 u) (fun (u:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154c x0 y0 u) m))). Time Defined. (* constant 2664 *) Definition l_e_st_eq_landau_n_rt_rp_satz154e : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_is x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_ec3_th10 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_5154_t4 x0 y0) (l_e_st_eq_landau_n_rt_rp_5154_t5 x0 y0) (fun (u:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154b x0 y0 u) (fun (u:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154a x0 y0 u) (fun (u:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154c x0 y0 u) i))). Time Defined. (* constant 2665 *) Definition l_e_st_eq_landau_n_rt_rp_satz154f : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_less x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_ec3_th12 (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_more x0 y0) (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_5154_t4 x0 y0) (l_e_st_eq_landau_n_rt_rp_5154_t5 x0 y0) (fun (u:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154b x0 y0 u) (fun (u:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154a x0 y0 u) (fun (u:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154c x0 y0 u) l))). Time Defined. (* constant 2666 *) Definition l_e_st_eq_landau_n_rt_rp_iii5_t2 : l_e_injective l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x). exact (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt y)) => l_e_st_eq_landau_n_rt_rp_satz154e x y t))). Time Defined. (* constant 2667 *) Definition l_e_st_eq_landau_n_rt_rp_isrterp : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154b x0 y0 i))). Time Defined. (* constant 2668 *) Definition l_e_st_eq_landau_n_rt_rp_isrtirp : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_is x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_e_st_eq_landau_n_rt_rp_satz154e x0 y0 i))). Time Defined. (* constant 2669 *) Definition l_e_st_eq_landau_n_rt_rp_rtofrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_rat)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_soft l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) l_e_st_eq_landau_n_rt_rp_iii5_t2 ksi rtksi)). Time Defined. (* constant 2670 *) Definition l_e_st_eq_landau_n_rt_rp_isrpert : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (rteta:l_e_st_eq_landau_n_rt_rp_ratrp eta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_rtofrp ksi rtksi) (l_e_st_eq_landau_n_rt_rp_rtofrp eta rteta)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (rteta:l_e_st_eq_landau_n_rt_rp_ratrp eta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isinv l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) l_e_st_eq_landau_n_rt_rp_iii5_t2 ksi rtksi eta rteta i))))). Time Defined. (* constant 2671 *) Definition l_e_st_eq_landau_n_rt_rp_isrpirt : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (rteta:l_e_st_eq_landau_n_rt_rp_ratrp eta), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_rtofrp ksi rtksi) (l_e_st_eq_landau_n_rt_rp_rtofrp eta rteta)), l_e_st_eq_landau_n_rt_rp_is ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (rteta:l_e_st_eq_landau_n_rt_rp_ratrp eta) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_rtofrp ksi rtksi) (l_e_st_eq_landau_n_rt_rp_rtofrp eta rteta)) => l_e_isinve l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) l_e_st_eq_landau_n_rt_rp_iii5_t2 ksi rtksi eta rteta i))))). Time Defined. (* constant 2672 *) Definition l_e_st_eq_landau_n_rt_rp_isrtrp1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_rp_rtofrp (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_isst1 l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) l_e_st_eq_landau_n_rt_rp_iii5_t2 x0). Time Defined. (* constant 2673 *) Definition l_e_st_eq_landau_n_rt_rp_isrtrp2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_rtofrp (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0)) x0). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_isst2 l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) l_e_st_eq_landau_n_rt_rp_iii5_t2 x0). Time Defined. (* constant 2674 *) Definition l_e_st_eq_landau_n_rt_rp_isrprt1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_rtofrp ksi rtksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_ists1 l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) l_e_st_eq_landau_n_rt_rp_iii5_t2 ksi rtksi)). Time Defined. (* constant 2675 *) Definition l_e_st_eq_landau_n_rt_rp_isrprt2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_rtofrp ksi rtksi)) ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_ists2 l_e_st_eq_landau_n_rt_rat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x) l_e_st_eq_landau_n_rt_rp_iii5_t2 ksi rtksi)). Time Defined. (* constant 2676 *) Definition l_e_st_eq_landau_n_rt_rp_isnterp : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (z:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt z) x y i))). Time Defined. (* constant 2677 *) Definition l_e_st_eq_landau_n_rt_rp_isntirp : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)), l_e_st_eq_landau_n_is x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)) => l_e_st_eq_landau_n_rt_isnirt x y (l_e_st_eq_landau_n_rt_rp_isrtirp (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y) i)))). Time Defined. (* constant 2678 *) Definition l_e_st_eq_landau_n_rt_rp_iii5_t3 : l_e_injective l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt x). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)) => l_e_st_eq_landau_n_rt_rp_isntirp x y t))). Time Defined. (* constant 2679 *) Definition l_e_st_eq_landau_n_rt_rp_ntofrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), l_e_st_eq_landau_n_nat)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => l_e_soft l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt x) l_e_st_eq_landau_n_rt_rp_iii5_t3 ksi ntksi)). Time Defined. (* constant 2680 *) Definition l_e_st_eq_landau_n_rt_rp_isrpent : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (nteta:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_ntofrp ksi ntksi) (l_e_st_eq_landau_n_rt_rp_ntofrp eta nteta)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (nteta:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isinv l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt x) l_e_st_eq_landau_n_rt_rp_iii5_t3 ksi ntksi eta nteta i))))). Time Defined. (* constant 2681 *) Definition l_e_st_eq_landau_n_rt_rp_isrpint : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (nteta:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_ntofrp ksi ntksi) (l_e_st_eq_landau_n_rt_rp_ntofrp eta nteta)), l_e_st_eq_landau_n_rt_rp_is ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (nteta:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_ntofrp ksi ntksi) (l_e_st_eq_landau_n_rt_rp_ntofrp eta nteta)) => l_e_isinve l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt x) l_e_st_eq_landau_n_rt_rp_iii5_t3 ksi ntksi eta nteta i))))). Time Defined. (* constant 2682 *) Definition l_e_st_eq_landau_n_rt_rp_isntrp1 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_rt_rp_ntofrp (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_isst1 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt y) l_e_st_eq_landau_n_rt_rp_iii5_t3 x). Time Defined. (* constant 2683 *) Definition l_e_st_eq_landau_n_rt_rp_isntrp2 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_ntofrp (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x)) x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_isst2 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt y) l_e_st_eq_landau_n_rt_rp_iii5_t3 x). Time Defined. (* constant 2684 *) Definition l_e_st_eq_landau_n_rt_rp_isrpnt1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_ntofrp ksi ntksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => l_e_ists1 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt x) l_e_st_eq_landau_n_rt_rp_iii5_t3 ksi ntksi)). Time Defined. (* constant 2685 *) Definition l_e_st_eq_landau_n_rt_rp_isrpnt2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_ntofrp ksi ntksi)) ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (ntksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => l_e_ists2 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_rpofnt x) l_e_st_eq_landau_n_rt_rp_iii5_t3 ksi ntksi)). Time Defined. (* constant 2686 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)), l_e_st_eq_landau_n_rt_less u0 x0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte x0 u0 lu))))))))). Time Defined. (* constant 2687 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t2 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)), l_e_st_eq_landau_n_rt_less v0 y0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte y0 v0 lv))))))))). Time Defined. (* constant 2688 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t3 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_pl u0 v0) (l_e_st_eq_landau_n_rt_pl x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)) => l_e_st_eq_landau_n_rt_satz98a u0 x0 v0 y0 (l_e_st_eq_landau_n_rt_rp_5155_t1 x0 y0 z0 lz u0 lu v0 lv i) (l_e_st_eq_landau_n_rt_rp_5155_t2 x0 y0 z0 lz u0 lu v0 lv i)))))))))). Time Defined. (* constant 2689 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t4 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)), l_e_st_eq_landau_n_rt_less z0 (l_e_st_eq_landau_n_rt_pl x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)) => l_e_st_eq_landau_n_rt_isless1 (l_e_st_eq_landau_n_rt_pl u0 v0) z0 (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_pl u0 v0) i) (l_e_st_eq_landau_n_rt_rp_5155_t3 x0 y0 z0 lz u0 lu v0 lv i)))))))))). Time Defined. (* constant 2690 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t5 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) z0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl u0 v0)) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt (l_e_st_eq_landau_n_rt_pl x0 y0) z0 (l_e_st_eq_landau_n_rt_rp_5155_t4 x0 y0 z0 lz u0 lu v0 lv i)))))))))). Time Defined. (* constant 2691 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t6 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) z0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => l_e_st_eq_landau_n_rt_rp_plapp (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) z0 lz (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_pl x y)) => l_e_st_eq_landau_n_rt_rp_5155_t5 x0 y0 z0 lz x t y u v))))))))). Time Defined. (* constant 2692 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t7 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0), l_e_st_eq_landau_n_rt_less u0 (l_e_st_eq_landau_n_rt_pl x0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte (l_e_st_eq_landau_n_rt_pl x0 y0) u0 lu)))). Time Defined. (* constant 2693 *) Definition l_e_st_eq_landau_n_rt_rp_5155_u01 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0), l_e_st_eq_landau_n_rt_rat)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0) => l_e_st_eq_landau_n_rt_ov u0 (l_e_st_eq_landau_n_rt_pl x0 y0))))). Time Defined. (* constant 2694 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t8 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu) (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_pl x0 y0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0) => l_e_st_eq_landau_n_rt_isless12 u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu) (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_satz110f u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_example1d (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_rp_5155_t7 x0 y0 u0 lu))))). Time Defined. (* constant 2695 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t9 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu) l_e_st_eq_landau_n_rt_1rt)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0) => l_e_st_eq_landau_n_rt_satz106c (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu) l_e_st_eq_landau_n_rt_1rt (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_rp_5155_t8 x0 y0 u0 lu))))). Time Defined. (* constant 2696 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t10 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu)) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0) => l_e_tris l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_pl x0 y0) (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu)) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu))) (l_e_st_eq_landau_n_rt_satz110d u0 (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_disttp1 x0 y0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu)))))). Time Defined. (* constant 2697 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t11 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 l_e_st_eq_landau_n_rt_1rt), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts x0 y0) x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 l_e_st_eq_landau_n_rt_1rt) => l_e_st_eq_landau_n_rt_isless12 (l_e_st_eq_landau_n_rt_ts y0 x0) (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt x0) x0 (l_e_st_eq_landau_n_rt_comts y0 x0) (l_e_st_eq_landau_n_rt_example1c x0) (l_e_st_eq_landau_n_rt_satz105c y0 l_e_st_eq_landau_n_rt_1rt x0 l)))). Time Defined. (* constant 2698 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t12 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 l_e_st_eq_landau_n_rt_1rt), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_ts x0 y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 l_e_st_eq_landau_n_rt_1rt) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt x0 (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_rp_5155_t11 x0 y0 l)))). Time Defined. (* constant 2699 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t13 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) u0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) u0) => l_e_st_eq_landau_n_rt_rp_lrtpl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) u0 (l_e_st_eq_landau_n_rt_ts x0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu)) (l_e_st_eq_landau_n_rt_rp_5155_t12 x0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu) (l_e_st_eq_landau_n_rt_rp_5155_t9 x0 y0 u0 lu)) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu)) (l_e_st_eq_landau_n_rt_rp_5155_t12 y0 (l_e_st_eq_landau_n_rt_rp_5155_u01 x0 y0 u0 lu) (l_e_st_eq_landau_n_rt_rp_5155_t9 x0 y0 u0 lu)) (l_e_st_eq_landau_n_rt_rp_5155_t10 x0 y0 u0 lu))))). Time Defined. (* constant 2700 *) Definition l_e_st_eq_landau_n_rt_rp_satz155a : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 y0)) x) => l_e_st_eq_landau_n_rt_rp_5155_t13 x0 y0 x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) x) => l_e_st_eq_landau_n_rt_rp_5155_t6 x0 y0 x t)))). Time Defined. (* constant 2701 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t14 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_satz101f x0 y0 m))). Time Defined. (* constant 2702 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t15 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn x0 y0 m)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn x0 y0 m)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_isrterp x0 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0) (l_e_st_eq_landau_n_rt_rp_5155_t14 x0 y0 m)) (l_e_st_eq_landau_n_rt_rp_satz155a (l_e_st_eq_landau_n_rt_mn x0 y0 m) y0)))). Time Defined. (* constant 2703 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t16 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn x0 y0 m))) (l_e_st_eq_landau_n_rt_rp_rpofrt x0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn x0 y0 m))) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn x0 y0 m)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn x0 y0 m))) (l_e_st_eq_landau_n_rt_rp_5155_t15 x0 y0 m)))). Time Defined. (* constant 2704 *) Definition l_e_st_eq_landau_n_rt_rp_satz155b : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_more x0 y0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn x0 y0 m)) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_satz154a x0 y0 m))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_more x0 y0) => l_e_st_eq_landau_n_rt_rp_satz140g (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn x0 y0 m)) (l_e_st_eq_landau_n_rt_rp_satz154a x0 y0 m) (l_e_st_eq_landau_n_rt_rp_5155_t16 x0 y0 m)))). Time Defined. (* constant 2705 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t17 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)), l_e_st_eq_landau_n_rt_less u0 x0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte x0 u0 lu))))))))). Time Defined. (* constant 2706 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t18 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)), l_e_st_eq_landau_n_rt_less v0 y0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte y0 v0 lv))))))))). Time Defined. (* constant 2707 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t19 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts u0 v0) (l_e_st_eq_landau_n_rt_ts x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)) => l_e_st_eq_landau_n_rt_satz107a u0 x0 v0 y0 (l_e_st_eq_landau_n_rt_rp_5155_t17 x0 y0 z0 lz u0 lu v0 lv i) (l_e_st_eq_landau_n_rt_rp_5155_t18 x0 y0 z0 lz u0 lu v0 lv i)))))))))). Time Defined. (* constant 2708 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t20 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)), l_e_st_eq_landau_n_rt_less z0 (l_e_st_eq_landau_n_rt_ts x0 y0)))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)) => l_e_st_eq_landau_n_rt_isless1 (l_e_st_eq_landau_n_rt_ts u0 v0) z0 (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_symis l_e_st_eq_landau_n_rt_rat z0 (l_e_st_eq_landau_n_rt_ts u0 v0) i) (l_e_st_eq_landau_n_rt_rp_5155_t19 x0 y0 z0 lz u0 lu v0 lv i)))))))))). Time Defined. (* constant 2709 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t21 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0), (forall (v0:l_e_st_eq_landau_n_rt_rat), (forall (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0), (forall (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) z0))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) u0) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => (fun (lv:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) v0) => (fun (i:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts u0 v0)) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt (l_e_st_eq_landau_n_rt_ts x0 y0) z0 (l_e_st_eq_landau_n_rt_rp_5155_t20 x0 y0 z0 lz u0 lu v0 lv i)))))))))). Time Defined. (* constant 2710 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t22 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) z0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) z0) => l_e_st_eq_landau_n_rt_rp_tsapp (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) z0 lz (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt y0) y) => (fun (v:l_e_st_eq_landau_n_rt_is z0 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_5155_t21 x0 y0 z0 lz x t y u v))))))))). Time Defined. (* constant 2711 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t23 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), l_e_st_eq_landau_n_rt_less u0 (l_e_st_eq_landau_n_rt_ts x0 y0))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => l_e_st_eq_landau_n_rt_rp_lrtrpofrte (l_e_st_eq_landau_n_rt_ts x0 y0) u0 lu)))). Time Defined. (* constant 2712 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t24 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), (forall (u1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))), l_e_st_eq_landau_n_rt_less u0 u1)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => (fun (u1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_ande1 (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0)) a)))))). Time Defined. (* constant 2713 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t25 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), (forall (u1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))), l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => (fun (u1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_ande2 (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0)) a)))))). Time Defined. (* constant 2714 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t26 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), (forall (u1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov u0 u1) u1) (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt u1))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => (fun (u1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_e_st_eq_landau_n_rt_isless12 u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov u0 u1) u1) u1 (l_e_st_eq_landau_n_rt_ts l_e_st_eq_landau_n_rt_1rt u1) (l_e_st_eq_landau_n_rt_satz110f u0 u1) (l_e_st_eq_landau_n_rt_example1d u1) (l_e_st_eq_landau_n_rt_rp_5155_t24 x0 y0 u0 lu u1 a))))))). Time Defined. (* constant 2715 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t27 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), (forall (u1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ov u0 u1) l_e_st_eq_landau_n_rt_1rt)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => (fun (u1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_e_st_eq_landau_n_rt_satz106c (l_e_st_eq_landau_n_rt_ov u0 u1) l_e_st_eq_landau_n_rt_1rt u1 (l_e_st_eq_landau_n_rt_rp_5155_t26 x0 y0 u0 lu u1 a))))))). Time Defined. (* constant 2716 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t28 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), (forall (u1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov u1 y0) y0) (l_e_st_eq_landau_n_rt_ts x0 y0))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => (fun (u1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_e_st_eq_landau_n_rt_isless1 u1 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov u1 y0) y0) (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_satz110f u1 y0) (l_e_st_eq_landau_n_rt_rp_5155_t25 x0 y0 u0 lu u1 a))))))). Time Defined. (* constant 2717 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t29 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), (forall (u1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ov u1 y0) x0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => (fun (u1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_e_st_eq_landau_n_rt_satz106c (l_e_st_eq_landau_n_rt_ov u1 y0) x0 y0 (l_e_st_eq_landau_n_rt_rp_5155_t28 x0 y0 u0 lu u1 a))))))). Time Defined. (* constant 2718 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t30 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), (forall (u1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))), l_e_st_eq_landau_n_rt_is u0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov u1 y0) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov u0 u1))))))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => (fun (u1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_e_tr3is l_e_st_eq_landau_n_rt_rat u0 (l_e_st_eq_landau_n_rt_ts u1 (l_e_st_eq_landau_n_rt_ov u0 u1)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov u1 y0) y0) (l_e_st_eq_landau_n_rt_ov u0 u1)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov u1 y0) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov u0 u1))) (l_e_st_eq_landau_n_rt_satz110d u0 u1) (l_e_st_eq_landau_n_rt_ists1 u1 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov u1 y0) y0) (l_e_st_eq_landau_n_rt_ov u0 u1) (l_e_st_eq_landau_n_rt_satz110f u1 y0)) (l_e_st_eq_landau_n_rt_assts1 (l_e_st_eq_landau_n_rt_ov u1 y0) y0 (l_e_st_eq_landau_n_rt_ov u0 u1)))))))). Time Defined. (* constant 2719 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t31 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), (forall (u1:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) u0)))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => (fun (u1:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_less u0 u1) (l_e_st_eq_landau_n_rt_less u1 (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_e_st_eq_landau_n_rt_rp_lrtts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) u0 (l_e_st_eq_landau_n_rt_ov u1 y0) (l_e_st_eq_landau_n_rt_rp_lrtrpofrt x0 (l_e_st_eq_landau_n_rt_ov u1 y0) (l_e_st_eq_landau_n_rt_rp_5155_t29 x0 y0 u0 lu u1 a)) (l_e_st_eq_landau_n_rt_ts y0 (l_e_st_eq_landau_n_rt_ov u0 u1)) (l_e_st_eq_landau_n_rt_rp_5155_t12 y0 (l_e_st_eq_landau_n_rt_ov u0 u1) (l_e_st_eq_landau_n_rt_rp_5155_t27 x0 y0 u0 lu u1 a)) (l_e_st_eq_landau_n_rt_rp_5155_t30 x0 y0 u0 lu u1 a))))))). Time Defined. (* constant 2720 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t32 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) u0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) u0) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_less u0 x) (l_e_st_eq_landau_n_rt_less x (l_e_st_eq_landau_n_rt_ts x0 y0))) (l_e_st_eq_landau_n_rt_satz91 u0 (l_e_st_eq_landau_n_rt_ts x0 y0) (l_e_st_eq_landau_n_rt_rp_5155_t23 x0 y0 u0 lu)) (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) u0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_and (l_e_st_eq_landau_n_rt_less u0 x) (l_e_st_eq_landau_n_rt_less x (l_e_st_eq_landau_n_rt_ts x0 y0))) => l_e_st_eq_landau_n_rt_rp_5155_t31 x0 y0 u0 lu x t)))))). Time Defined. (* constant 2721 *) Definition l_e_st_eq_landau_n_rt_rp_satz155c : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_isi1 (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x0 y0)) x) => l_e_st_eq_landau_n_rt_rp_5155_t32 x0 y0 x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) x) => l_e_st_eq_landau_n_rt_rp_5155_t22 x0 y0 x t)))). Time Defined. (* constant 2722 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t33 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov x0 y0) y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz110f x0 y0)). Time Defined. (* constant 2723 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t34 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ov x0 y0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov x0 y0) y0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ov x0 y0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_isrterp x0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_ov x0 y0) y0) (l_e_st_eq_landau_n_rt_rp_5155_t33 x0 y0)) (l_e_st_eq_landau_n_rt_rp_satz155c (l_e_st_eq_landau_n_rt_ov x0 y0) y0))). Time Defined. (* constant 2724 *) Definition l_e_st_eq_landau_n_rt_rp_5155_t35 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ov x0 y0))) (l_e_st_eq_landau_n_rt_rp_rpofrt x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ov x0 y0))) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ov x0 y0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ov x0 y0))) (l_e_st_eq_landau_n_rt_rp_5155_t34 x0 y0))). Time Defined. (* constant 2725 *) Definition l_e_st_eq_landau_n_rt_rp_satz155d : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ov x0 y0)) (l_e_st_eq_landau_n_rt_rp_ov (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_satz153g (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ov x0 y0)) (l_e_st_eq_landau_n_rt_rp_5155_t35 x0 y0))). Time Defined. (* constant 2726 *) Definition l_e_st_eq_landau_n_rt_rp_satz155e : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)) (l_e_st_eq_landau_n_rt_rp_isrterp (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_satz112h x y))) (l_e_st_eq_landau_n_rt_rp_satz155a (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)))). Time Defined. (* constant 2727 *) Definition l_e_st_eq_landau_n_rt_rp_satz155f : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)) (l_e_st_eq_landau_n_rt_rp_isrterp (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_satz112j x y))) (l_e_st_eq_landau_n_rt_rp_satz155c (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)))). Time Defined. (* constant 2728 *) Definition l_e_st_eq_landau_n_rt_rp_nt_natt : Type. exact (l_e_ot l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t)). Time Defined. (* constant 2729 *) Definition l_e_st_eq_landau_n_rt_rp_nt_nttofrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), l_e_st_eq_landau_n_rt_rp_nt_natt)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => l_e_out l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) ksi nksi)). Time Defined. (* constant 2730 *) Definition l_e_st_eq_landau_n_rt_rp_nt_is : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_is l_e_st_eq_landau_n_rt_rp_nt_natt xt yt)). Time Defined. (* constant 2731 *) Definition l_e_st_eq_landau_n_rt_rp_nt_nis : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_not (l_e_st_eq_landau_n_rt_rp_nt_is xt yt))). Time Defined. (* constant 2732 *) Definition l_e_st_eq_landau_n_rt_rp_nt_all : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)) => l_all l_e_st_eq_landau_n_rt_rp_nt_natt p). Time Defined. (* constant 2733 *) Definition l_e_st_eq_landau_n_rt_rp_nt_some : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)) => l_some l_e_st_eq_landau_n_rt_rp_nt_natt p). Time Defined. (* constant 2734 *) Definition l_e_st_eq_landau_n_rt_rp_nt_one : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)) => l_e_one l_e_st_eq_landau_n_rt_rp_nt_natt p). Time Defined. (* constant 2735 *) Definition l_e_st_eq_landau_n_rt_rp_nt_in : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), Prop)). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_esti l_e_st_eq_landau_n_rt_rp_nt_natt xt st)). Time Defined. (* constant 2736 *) Definition l_e_st_eq_landau_n_rt_rp_nt_rpofntt : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_cut). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_in l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) xt). Time Defined. (* constant 2737 *) Definition l_e_st_eq_landau_n_rt_rp_nt_natrpi : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt)). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_inp l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) xt). Time Defined. (* constant 2738 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isrpentt : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (neta:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_nttofrp ksi nksi) (l_e_st_eq_landau_n_rt_rp_nt_nttofrp eta neta)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (neta:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isouti l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) ksi nksi eta neta i))))). Time Defined. (* constant 2739 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isrpintt : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (neta:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (i:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_nttofrp ksi nksi) (l_e_st_eq_landau_n_rt_rp_nt_nttofrp eta neta)), l_e_st_eq_landau_n_rt_rp_is ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (neta:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (i:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_nttofrp ksi nksi) (l_e_st_eq_landau_n_rt_rp_nt_nttofrp eta neta)) => l_e_isoute l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) ksi nksi eta neta i))))). Time Defined. (* constant 2740 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isntterp : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_rp_nt_is xt yt), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_rpofntt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_rp_nt_is xt yt) => l_e_isini l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) xt yt i))). Time Defined. (* constant 2741 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isnttirp : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_rpofntt yt)), l_e_st_eq_landau_n_rt_rp_nt_is xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_rpofntt yt)) => l_e_isine l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) xt yt i))). Time Defined. (* constant 2742 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isrpntt1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nksi:l_e_st_eq_landau_n_rt_rp_natrp ksi), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_nt_rpofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofrp ksi nksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nksi:l_e_st_eq_landau_n_rt_rp_natrp ksi) => l_e_isinout l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) ksi nksi)). Time Defined. (* constant 2743 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isnttrp1 : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_is xt (l_e_st_eq_landau_n_rt_rp_nt_nttofrp (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi xt))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_isoutin l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) xt). Time Defined. (* constant 2744 *) Definition l_e_st_eq_landau_n_rt_rp_nt_nttofnt : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_nt_natt). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_nt_nttofrp (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x)). Time Defined. (* constant 2745 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isntentt : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x) (l_e_st_eq_landau_n_rt_rp_nt_nttofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_st_eq_landau_n_rt_rp_nt_isrpentt (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x) (l_e_st_eq_landau_n_rt_rp_rpofnt y) (l_e_st_eq_landau_n_rt_rp_natrpi y) (l_e_st_eq_landau_n_rt_rp_isnterp x y i)))). Time Defined. (* constant 2746 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isntintt : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x) (l_e_st_eq_landau_n_rt_rp_nt_nttofnt y)), l_e_st_eq_landau_n_is x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x) (l_e_st_eq_landau_n_rt_rp_nt_nttofnt y)) => l_e_st_eq_landau_n_rt_rp_isntirp x y (l_e_st_eq_landau_n_rt_rp_nt_isrpintt (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x) (l_e_st_eq_landau_n_rt_rp_rpofnt y) (l_e_st_eq_landau_n_rt_rp_natrpi y) i)))). Time Defined. (* constant 2747 *) Definition l_e_st_eq_landau_n_rt_rp_nt_ntofntt : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_nat). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_ntofrp (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi xt)). Time Defined. (* constant 2748 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isnttent : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_rp_nt_is xt yt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_ntofntt yt)))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_rp_nt_is xt yt) => l_e_st_eq_landau_n_rt_rp_isrpent (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi xt) (l_e_st_eq_landau_n_rt_rp_nt_rpofntt yt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi yt) (l_e_st_eq_landau_n_rt_rp_nt_isntterp xt yt i)))). Time Defined. (* constant 2749 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isnttint : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_ntofntt yt)), l_e_st_eq_landau_n_rt_rp_nt_is xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_ntofntt yt)) => l_e_st_eq_landau_n_rt_rp_nt_isnttirp xt yt (l_e_st_eq_landau_n_rt_rp_isrpint (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi xt) (l_e_st_eq_landau_n_rt_rp_nt_rpofntt yt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi yt) i)))). Time Defined. (* constant 2750 *) Definition l_e_st_eq_landau_n_rt_rp_nt_iii5_t5 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_nt_rpofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_nt_isrpntt1 (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x)). Time Defined. (* constant 2751 *) Definition l_e_st_eq_landau_n_rt_rp_nt_iii5_t6 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_ntofrp (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x)) (l_e_st_eq_landau_n_rt_rp_nt_ntofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_isrpent (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x) (l_e_st_eq_landau_n_rt_rp_nt_rpofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x)) (l_e_st_eq_landau_n_rt_rp_nt_natrpi (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x)) (l_e_st_eq_landau_n_rt_rp_nt_iii5_t5 x)). Time Defined. (* constant 2752 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isntntt1 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_rt_rp_nt_ntofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_rt_rp_ntofrp (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_natrpi x)) (l_e_st_eq_landau_n_rt_rp_nt_ntofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x)) (l_e_st_eq_landau_n_rt_rp_isntrp1 x) (l_e_st_eq_landau_n_rt_rp_nt_iii5_t6 x)). Time Defined. (* constant 2753 *) Definition l_e_st_eq_landau_n_rt_rp_nt_iii5_t7 : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_isrpnt1 (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi xt)). Time Defined. (* constant 2754 *) Definition l_e_st_eq_landau_n_rt_rp_nt_iii5_t8 : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_nttofrp (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi xt)) (l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_nt_isrpentt (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi xt) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) (l_e_st_eq_landau_n_rt_rp_natrpi (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) (l_e_st_eq_landau_n_rt_rp_nt_iii5_t7 xt)). Time Defined. (* constant 2755 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isnttnt1 : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_is xt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_tris l_e_st_eq_landau_n_rt_rp_nt_natt xt (l_e_st_eq_landau_n_rt_rp_nt_nttofrp (l_e_st_eq_landau_n_rt_rp_nt_rpofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_natrpi xt)) (l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) (l_e_st_eq_landau_n_rt_rp_nt_isnttrp1 xt) (l_e_st_eq_landau_n_rt_rp_nt_iii5_t8 xt)). Time Defined. (* constant 2756 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isntntt2 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_nt_ntofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x)) x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_rt_rp_nt_ntofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x)) (l_e_st_eq_landau_n_rt_rp_nt_isntntt1 x)). Time Defined. (* constant 2757 *) Definition l_e_st_eq_landau_n_rt_rp_nt_isnttnt2 : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) xt). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_symis l_e_st_eq_landau_n_rt_rp_nt_natt xt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) (l_e_st_eq_landau_n_rt_rp_nt_isnttnt1 xt)). Time Defined. (* constant 2758 *) Definition l_e_st_eq_landau_n_rt_rp_nt_1t : l_e_st_eq_landau_n_rt_rp_nt_natt. exact (l_e_st_eq_landau_n_rt_rp_nt_nttofnt l_e_st_eq_landau_n_1). Time Defined. (* constant 2759 *) Definition l_e_st_eq_landau_n_rt_rp_nt_suct : (forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_natt). exact (fun (x:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_nt_ntofntt x))). Time Defined. (* constant 2760 *) Definition l_e_st_eq_landau_n_rt_rp_nt_5156_t1 : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (j:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct xt) l_e_st_eq_landau_n_rt_rp_nt_1t), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) l_e_st_eq_landau_n_1)). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (j:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct xt) l_e_st_eq_landau_n_rt_rp_nt_1t) => l_e_st_eq_landau_n_rt_rp_nt_isntintt (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) l_e_st_eq_landau_n_1 j)). Time Defined. (* constant 2761 *) Definition l_e_st_eq_landau_n_rt_rp_nt_satz156a : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_nis (l_e_st_eq_landau_n_rt_rp_nt_suct xt) l_e_st_eq_landau_n_rt_rp_nt_1t). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct xt) l_e_st_eq_landau_n_rt_rp_nt_1t) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ax3 (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) (fun (t:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct xt) l_e_st_eq_landau_n_rt_rp_nt_1t) => l_e_st_eq_landau_n_rt_rp_nt_5156_t1 xt t)). Time Defined. (* constant 2762 *) Definition l_e_st_eq_landau_n_rt_rp_nt_5156_t2 : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct xt) (l_e_st_eq_landau_n_rt_rp_nt_suct yt)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_nt_ntofntt yt))))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct xt) (l_e_st_eq_landau_n_rt_rp_nt_suct yt)) => l_e_st_eq_landau_n_rt_rp_nt_isntintt (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_nt_ntofntt yt)) i))). Time Defined. (* constant 2763 *) Definition l_e_st_eq_landau_n_rt_rp_nt_satz156b : (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (yt:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (i:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct xt) (l_e_st_eq_landau_n_rt_rp_nt_suct yt)), l_e_st_eq_landau_n_rt_rp_nt_is xt yt))). exact (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (yt:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (i:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct xt) (l_e_st_eq_landau_n_rt_rp_nt_suct yt)) => l_e_st_eq_landau_n_rt_rp_nt_isnttint xt yt (l_e_st_eq_landau_n_ax4 (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt) (l_e_st_eq_landau_n_rt_rp_nt_ntofntt yt) (l_e_st_eq_landau_n_rt_rp_nt_5156_t2 xt yt i))))). Time Defined. (* constant 2764 *) Definition l_e_st_eq_landau_n_rt_rp_nt_cond1 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), Prop). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_nt_in l_e_st_eq_landau_n_rt_rp_nt_1t st). Time Defined. (* constant 2765 *) Definition l_e_st_eq_landau_n_rt_rp_nt_cond2 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), Prop). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_nt_all (fun (x:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_imp (l_e_st_eq_landau_n_rt_rp_nt_in x st) (l_e_st_eq_landau_n_rt_rp_nt_in (l_e_st_eq_landau_n_rt_rp_nt_suct x) st))). Time Defined. (* constant 2766 *) Definition l_e_st_eq_landau_n_rt_rp_nt_5156_prop1 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st), (forall (x:l_e_st_eq_landau_n_nat), Prop)))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_nt_in (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x) st)))). Time Defined. (* constant 2767 *) Definition l_e_st_eq_landau_n_rt_rp_nt_5156_t3 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_nt_5156_prop1 st c1 c2 x), l_e_st_eq_landau_n_rt_rp_nt_in (l_e_st_eq_landau_n_rt_rp_nt_suct (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x)) st))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_nt_5156_prop1 st c1 c2 x) => c2 (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x) p))))). Time Defined. (* constant 2768 *) Definition l_e_st_eq_landau_n_rt_rp_nt_5156_t4 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_nt_5156_prop1 st c1 c2 x), l_e_st_eq_landau_n_rt_rp_nt_5156_prop1 st c1 c2 (l_e_st_eq_landau_n_suc x)))))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_nt_5156_prop1 st c1 c2 x) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_nt_in (l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_suc t)) st) (l_e_st_eq_landau_n_rt_rp_nt_ntofntt (l_e_st_eq_landau_n_rt_rp_nt_nttofnt x)) x (l_e_st_eq_landau_n_rt_rp_nt_5156_t3 st c1 c2 x p) (l_e_st_eq_landau_n_rt_rp_nt_isntntt2 x)))))). Time Defined. (* constant 2769 *) Definition l_e_st_eq_landau_n_rt_rp_nt_5156_t5 : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_in (l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) st)))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_nt_5156_prop1 st c1 c2 t) c1 (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_rp_nt_5156_prop1 st c1 c2 t) => l_e_st_eq_landau_n_rt_rp_nt_5156_t4 st c1 c2 t u)) (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt))))). Time Defined. (* constant 2770 *) Definition l_e_st_eq_landau_n_rt_rp_nt_satz156c : (forall (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st), (forall (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st), (forall (xt:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_in xt st)))). exact (fun (st:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_nt_cond1 st) => (fun (c2:l_e_st_eq_landau_n_rt_rp_nt_cond2 st) => (fun (xt:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_isp l_e_st_eq_landau_n_rt_rp_nt_natt (fun (t:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_nt_in t st) (l_e_st_eq_landau_n_rt_rp_nt_nttofnt (l_e_st_eq_landau_n_rt_rp_nt_ntofntt xt)) xt (l_e_st_eq_landau_n_rt_rp_nt_5156_t5 st c1 c2 xt) (l_e_st_eq_landau_n_rt_rp_nt_isnttnt2 xt))))). Time Defined. (* constant 2771 *) Definition l_e_st_eq_landau_n_rt_rp_nt_ax3t : (forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_nis (l_e_st_eq_landau_n_rt_rp_nt_suct x) l_e_st_eq_landau_n_rt_rp_nt_1t). exact (fun (x:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_nt_satz156a x). Time Defined. (* constant 2772 *) Definition l_e_st_eq_landau_n_rt_rp_nt_ax4t : (forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (y:l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (u:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct x) (l_e_st_eq_landau_n_rt_rp_nt_suct y)), l_e_st_eq_landau_n_rt_rp_nt_is x y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (y:l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (u:l_e_st_eq_landau_n_rt_rp_nt_is (l_e_st_eq_landau_n_rt_rp_nt_suct x) (l_e_st_eq_landau_n_rt_rp_nt_suct y)) => l_e_st_eq_landau_n_rt_rp_nt_satz156b x y u))). Time Defined. (* constant 2773 *) Definition l_e_st_eq_landau_n_rt_rp_nt_ax5t : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt), (forall (u:l_e_st_eq_landau_n_rt_rp_nt_cond1 s), (forall (v:l_e_st_eq_landau_n_rt_rp_nt_cond2 s), (forall (x:l_e_st_eq_landau_n_rt_rp_nt_natt), l_e_st_eq_landau_n_rt_rp_nt_in x s)))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_nt_natt) => (fun (u:l_e_st_eq_landau_n_rt_rp_nt_cond1 s) => (fun (v:l_e_st_eq_landau_n_rt_rp_nt_cond2 s) => (fun (x:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_nt_satz156c s u v x)))). Time Defined. (* constant 2774 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_ratt : Type. exact (l_e_ot l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t)). Time Defined. (* constant 2775 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_rttofrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_rp_rtt_ratt)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_out l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) ksi rtksi)). Time Defined. (* constant 2776 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_is : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), Prop)). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_e_is l_e_st_eq_landau_n_rt_rp_rtt_ratt x0t y0t)). Time Defined. (* constant 2777 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_nis : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), Prop)). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_not (l_e_st_eq_landau_n_rt_rp_rtt_is x0t y0t))). Time Defined. (* constant 2778 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_all : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_rtt_ratt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_rtt_ratt), Prop)) => l_all l_e_st_eq_landau_n_rt_rp_rtt_ratt p). Time Defined. (* constant 2779 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_some : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_rtt_ratt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_rtt_ratt), Prop)) => l_some l_e_st_eq_landau_n_rt_rp_rtt_ratt p). Time Defined. (* constant 2780 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_one : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_rtt_ratt), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_rtt_ratt), Prop)) => l_e_one l_e_st_eq_landau_n_rt_rp_rtt_ratt p). Time Defined. (* constant 2781 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), l_e_st_eq_landau_n_rt_cut). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_e_in l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) x0t). Time Defined. (* constant 2782 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_ratrpi : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t)). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_e_inp l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) x0t). Time Defined. (* constant 2783 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrpertt : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (rteta:l_e_st_eq_landau_n_rt_rp_ratrp eta), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_rtt_is (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp ksi rtksi) (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp eta rteta)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (rteta:l_e_st_eq_landau_n_rt_rp_ratrp eta) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isouti l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) ksi rtksi eta rteta i))))). Time Defined. (* constant 2784 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrpirtt : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (rteta:l_e_st_eq_landau_n_rt_rp_ratrp eta), (forall (i:l_e_st_eq_landau_n_rt_rp_rtt_is (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp ksi rtksi) (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp eta rteta)), l_e_st_eq_landau_n_rt_rp_is ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (rteta:l_e_st_eq_landau_n_rt_rp_ratrp eta) => (fun (i:l_e_st_eq_landau_n_rt_rp_rtt_is (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp ksi rtksi) (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp eta rteta)) => l_e_isoute l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) ksi rtksi eta rteta i))))). Time Defined. (* constant 2785 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrtterp : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (i:l_e_st_eq_landau_n_rt_rp_rtt_is x0t y0t), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt y0t)))). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (i:l_e_st_eq_landau_n_rt_rp_rtt_is x0t y0t) => l_e_isini l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) x0t y0t i))). Time Defined. (* constant 2786 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrttirp : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt y0t)), l_e_st_eq_landau_n_rt_rp_rtt_is x0t y0t))). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt y0t)) => l_e_isine l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) x0t y0t i))). Time Defined. (* constant 2787 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrprtt1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp ksi rtksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_isinout l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) ksi rtksi)). Time Defined. (* constant 2788 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrttrp1 : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), l_e_st_eq_landau_n_rt_rp_rtt_is x0t (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi x0t))). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_e_isoutin l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) x0t). Time Defined. (* constant 2789 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_rttofrt : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_rtt_ratt). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rtt_rttofrp (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0)). Time Defined. (* constant 2790 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrtertt : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is x0 y0), l_e_st_eq_landau_n_rt_rp_rtt_is (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0) (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_rp_rtt_isrpertt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_ratrpi y0) (l_e_st_eq_landau_n_rt_rp_isrterp x0 y0 i)))). Time Defined. (* constant 2791 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrtirtt : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_rp_rtt_is (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0) (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt y0)), l_e_st_eq_landau_n_rt_is x0 y0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_rp_rtt_is (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0) (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt y0)) => l_e_st_eq_landau_n_rt_rp_isrtirp x0 y0 (l_e_st_eq_landau_n_rt_rp_rtt_isrpirtt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_ratrpi y0) i)))). Time Defined. (* constant 2792 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), l_e_st_eq_landau_n_rt_rat). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_e_st_eq_landau_n_rt_rp_rtofrp (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi x0t)). Time Defined. (* constant 2793 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrttert : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (i:l_e_st_eq_landau_n_rt_rp_rtt_is x0t y0t), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt y0t)))). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (i:l_e_st_eq_landau_n_rt_rp_rtt_is x0t y0t) => l_e_st_eq_landau_n_rt_rp_isrpert (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi x0t) (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt y0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi y0t) (l_e_st_eq_landau_n_rt_rp_rtt_isrtterp x0t y0t i)))). Time Defined. (* constant 2794 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrttirt : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt y0t)), l_e_st_eq_landau_n_rt_rp_rtt_is x0t y0t))). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (y0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt y0t)) => l_e_st_eq_landau_n_rt_rp_rtt_isrttirp x0t y0t (l_e_st_eq_landau_n_rt_rp_isrpirt (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi x0t) (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt y0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi y0t) i)))). Time Defined. (* constant 2795 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_iii5_t9 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rtt_isrprtt1 (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0)). Time Defined. (* constant 2796 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_iii5_t10 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_rp_rtofrp (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0)) (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_isrpert (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0) (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0)) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0)) (l_e_st_eq_landau_n_rt_rp_rtt_iii5_t9 x0)). Time Defined. (* constant 2797 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrtrtt1 : (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_is x0 (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_tris l_e_st_eq_landau_n_rt_rat x0 (l_e_st_eq_landau_n_rt_rp_rtofrp (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_ratrpi x0)) (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt x0)) (l_e_st_eq_landau_n_rt_rp_isrtrp1 x0) (l_e_st_eq_landau_n_rt_rp_rtt_iii5_t10 x0)). Time Defined. (* constant 2798 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_iii5_t11 : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t))). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_e_st_eq_landau_n_rt_rp_isrprt1 (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi x0t)). Time Defined. (* constant 2799 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_iii5_t12 : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), l_e_st_eq_landau_n_rt_rp_rtt_is (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi x0t)) (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t))). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_e_st_eq_landau_n_rt_rp_rtt_isrpertt (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi x0t) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t)) (l_e_st_eq_landau_n_rt_rp_ratrpi (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t)) (l_e_st_eq_landau_n_rt_rp_rtt_iii5_t11 x0t)). Time Defined. (* constant 2800 *) Definition l_e_st_eq_landau_n_rt_rp_rtt_isrttrt1 : (forall (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt), l_e_st_eq_landau_n_rt_rp_rtt_is x0t (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t))). exact (fun (x0t:l_e_st_eq_landau_n_rt_rp_rtt_ratt) => l_e_tris l_e_st_eq_landau_n_rt_rp_rtt_ratt x0t (l_e_st_eq_landau_n_rt_rp_rtt_rttofrp (l_e_st_eq_landau_n_rt_rp_rtt_rpofrtt x0t) (l_e_st_eq_landau_n_rt_rp_rtt_ratrpi x0t)) (l_e_st_eq_landau_n_rt_rp_rtt_rttofrt (l_e_st_eq_landau_n_rt_rp_rtt_rtofrtt x0t)) (l_e_st_eq_landau_n_rt_rp_rtt_isrttrp1 x0t) (l_e_st_eq_landau_n_rt_rp_rtt_iii5_t12 x0t)). Time Defined. (* constant 2801 *) Definition l_e_st_eq_landau_n_rt_rp_example2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi (l_e_st_eq_landau_n_rt_rp_ov l_e_st_eq_landau_n_rt_rp_1rp ksi)) l_e_st_eq_landau_n_rt_rp_1rp). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_satz153c l_e_st_eq_landau_n_rt_rp_1rp ksi). Time Defined. (* constant 2802 *) Definition l_e_st_eq_landau_n_rt_rp_5157_x01 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_rat)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_st_eq_landau_n_rt_rp_rtofrp ksi rtksi)). Time Defined. (* constant 2803 *) Definition l_e_st_eq_landau_n_rt_rp_5157_s1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), l_e_st_set l_e_st_eq_landau_n_rt_rat). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)). Time Defined. (* constant 2804 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)), l_e_st_eq_landau_n_rt_urt ksi y0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x) y0 i)))). Time Defined. (* constant 2805 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi)) y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0 (l_e_st_eq_landau_n_rt_satz82 (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0 m)))))). Time Defined. (* constant 2806 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)), (forall (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0), l_e_st_eq_landau_n_rt_lrt ksi y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) => (fun (m:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_lrt x y0) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi)) ksi (l_e_st_eq_landau_n_rt_rp_5157_t2 ksi rtksi y0 i m) (l_e_st_eq_landau_n_rt_rp_isrprt2 ksi rtksi)))))). Time Defined. (* constant 2807 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)), l_not (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) => l_imp_th3 (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0) (l_e_st_eq_landau_n_rt_lrt ksi y0) (l_e_st_eq_landau_n_rt_rp_5157_t1 ksi rtksi y0 i) (fun (t:l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0) => l_e_st_eq_landau_n_rt_rp_5157_t3 ksi rtksi y0 i t))))). Time Defined. (* constant 2808 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)), l_e_st_eq_landau_n_rt_lessis (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) => l_e_st_eq_landau_n_rt_satz81e (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) y0 (l_e_st_eq_landau_n_rt_rp_5157_t4 ksi rtksi y0 i))))). Time Defined. (* constant 2809 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_lb (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) => l_e_st_eq_landau_n_rt_rp_5157_t5 ksi rtksi x t)))). Time Defined. (* constant 2810 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi)) (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_st_eq_landau_n_rt_rp_urtrpofrt (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) (l_e_st_eq_landau_n_rt_moreisi2 (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) (l_e_refis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi))))). Time Defined. (* constant 2811 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_urt ksi (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_urt x (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi)) ksi (l_e_st_eq_landau_n_rt_rp_5157_t7 ksi rtksi) (l_e_st_eq_landau_n_rt_rp_isrprt2 ksi rtksi))). Time Defined. (* constant 2812 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x) (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) (l_e_st_eq_landau_n_rt_rp_5157_t8 ksi rtksi))). Time Defined. (* constant 2813 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_andi (l_e_st_eq_landau_n_rt_lb (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi)) (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) (l_e_st_eq_landau_n_rt_rp_5157_t6 ksi rtksi) (l_e_st_eq_landau_n_rt_rp_5157_t9 ksi rtksi))). Time Defined. (* constant 2814 *) Definition l_e_st_eq_landau_n_rt_rp_satz157a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) (l_e_st_eq_landau_n_rt_rp_rtofrp ksi rtksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_e_st_eq_landau_n_rt_rp_5157_t10 ksi rtksi)). Time Defined. (* constant 2815 *) Definition l_e_st_eq_landau_n_rt_rp_satz157b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi y)) x))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (rtksi:l_e_st_eq_landau_n_rt_rp_ratrp ksi) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) x) (l_e_st_eq_landau_n_rt_rp_5157_x01 ksi rtksi) (l_e_st_eq_landau_n_rt_rp_5157_t10 ksi rtksi))). Time Defined. (* constant 2816 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), l_e_st_eq_landau_n_rt_lb (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) x0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => l_ande1 (l_e_st_eq_landau_n_rt_lb (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) x0) (l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) m))). Time Defined. (* constant 2817 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => l_ande2 (l_e_st_eq_landau_n_rt_lb (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) x0) (l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)) m))). Time Defined. (* constant 2818 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), l_e_st_eq_landau_n_rt_urt ksi x0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x) x0 (l_e_st_eq_landau_n_rt_rp_5157_t12 ksi x0 m)))). Time Defined. (* constant 2819 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), l_e_st_eq_landau_n_rt_less y0 x0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => l_e_st_eq_landau_n_rt_cutapp2a ksi y0 ly x0 (l_e_st_eq_landau_n_rt_rp_5157_t13 ksi x0 m)))))). Time Defined. (* constant 2820 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt ksi y0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt ksi y0) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt x0 y0 (l_e_st_eq_landau_n_rt_rp_5157_t14 ksi x0 m y0 ly)))))). Time Defined. (* constant 2821 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x) y0 uy))))). Time Defined. (* constant 2822 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_rt_moreis y0 x0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_e_st_eq_landau_n_rt_satz85 x0 y0 (l_e_st_eq_landau_n_rt_rp_5157_t11 ksi x0 m y0 (l_e_st_eq_landau_n_rt_rp_5157_t17 ksi x0 m y0 uy))))))). Time Defined. (* constant 2823 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt ksi y0), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt ksi y0) => l_e_st_eq_landau_n_rt_rp_urtrpofrt x0 y0 (l_e_st_eq_landau_n_rt_rp_5157_t18 ksi x0 m y0 uy)))))). Time Defined. (* constant 2824 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_imp (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) y0) (l_e_st_eq_landau_n_rt_lrt ksi y0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_cp (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) y0) (l_e_st_eq_landau_n_rt_lrt ksi y0) (fun (t:l_e_st_eq_landau_n_rt_urt ksi y0) => l_e_st_eq_landau_n_rt_rp_5157_t19 ksi x0 m y0 t))))). Time Defined. (* constant 2825 *) Definition l_e_st_eq_landau_n_rt_rp_satz157c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)) x0) => l_e_st_eq_landau_n_rt_rp_isi1 ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x) => l_e_st_eq_landau_n_rt_rp_5157_t15 ksi x0 m x t)) (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5157_t20 ksi x0 m x)))). Time Defined. (* constant 2826 *) Definition l_e_st_eq_landau_n_rt_rp_5157_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi y)) x)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) x0), l_e_st_eq_landau_n_rt_rp_ratrp ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi y)) x)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) x0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) x0 (l_e_st_eq_landau_n_rt_rp_satz157c ksi x0 m))))). Time Defined. (* constant 2827 *) Definition l_e_st_eq_landau_n_rt_rp_satz157d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi y)) x)), l_e_st_eq_landau_n_rt_rp_ratrp ksi)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_min (l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi y)) x)) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) x) s (l_e_st_eq_landau_n_rt_rp_ratrp ksi) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5157_s1 ksi) x) => l_e_st_eq_landau_n_rt_rp_5157_t21 ksi s x t)))). Time Defined. (* constant 2828 *) Definition l_e_st_eq_landau_n_rt_rp_5158_xr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x0)). Time Defined. (* constant 2829 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) x0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_eq_landau_n_rt_rp_urtrpofrt x0 x0 (l_e_st_eq_landau_n_rt_moreisi2 x0 x0 (l_e_refis l_e_st_eq_landau_n_rt_rat x0))))). Time Defined. (* constant 2830 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_and (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) x0) (l_e_st_eq_landau_n_rt_lrt ksi x0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_andi (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) x0) (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_rp_5158_t1 ksi x0 lx) lx))). Time Defined. (* constant 2831 *) Definition l_e_st_eq_landau_n_rt_rp_satz158a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt ksi x0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_rpofrt x0) x) (l_e_st_eq_landau_n_rt_lrt ksi x)) x0 (l_e_st_eq_landau_n_rt_rp_5158_t2 ksi x0 lx)))). Time Defined. (* constant 2832 *) Definition l_e_st_eq_landau_n_rt_rp_5158_s1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_set l_e_st_eq_landau_n_rt_rat))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x)))). Time Defined. (* constant 2833 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0) => l_e_symis l_e_st_eq_landau_n_rt_cut ksi (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) (l_e_st_eq_landau_n_rt_rp_satz157c ksi x0 m))))). Time Defined. (* constant 2834 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (m:l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (m:l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0) => l_e_st_eq_landau_n_rt_rp_moreisi2 (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi (l_e_st_eq_landau_n_rt_rp_5158_t3 ksi x0 ux m))))). Time Defined. (* constant 2835 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x) x0 ux)))). Time Defined. (* constant 2836 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), l_not (l_e_st_eq_landau_n_rt_lb (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => l_and_th4 (l_e_st_eq_landau_n_rt_lb (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0) (l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) n (l_e_st_eq_landau_n_rt_rp_5158_t5 ksi x0 ux n))))). Time Defined. (* constant 2837 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_not (l_imp (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 x))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => l_some_th1 l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_imp (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 x)) (l_e_st_eq_landau_n_rt_rp_5158_t6 ksi x0 ux n))))). Time Defined. (* constant 2838 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))), l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))) => l_imp_th5 (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0) o)))))). Time Defined. (* constant 2839 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))), l_e_st_eq_landau_n_rt_urt ksi y0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_urt ksi x) y0 (l_e_st_eq_landau_n_rt_rp_5158_t8 ksi x0 ux n y0 o))))))). Time Defined. (* constant 2840 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))), l_not (l_e_st_eq_landau_n_rt_lessis x0 y0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))) => l_imp_th6 (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0) o)))))). Time Defined. (* constant 2841 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))), l_e_st_eq_landau_n_rt_less y0 x0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))) => l_e_st_eq_landau_n_rt_satz82 x0 y0 (l_e_st_eq_landau_n_rt_satz81k x0 y0 (l_e_st_eq_landau_n_rt_rp_5158_t10 ksi x0 ux n y0 o)))))))). Time Defined. (* constant 2842 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) y0)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))) => l_e_st_eq_landau_n_rt_rp_lrtrpofrt x0 y0 (l_e_st_eq_landau_n_rt_rp_5158_t11 ksi x0 ux n y0 o))))))). Time Defined. (* constant 2843 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))), l_and (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) y0) (l_e_st_eq_landau_n_rt_urt ksi y0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))) => l_andi (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) y0) (l_e_st_eq_landau_n_rt_urt ksi y0) (l_e_st_eq_landau_n_rt_rp_5158_t12 ksi x0 ux n y0 o) (l_e_st_eq_landau_n_rt_rp_5158_t9 ksi x0 ux n y0 o))))))). Time Defined. (* constant 2844 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (o:l_not (l_imp (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 y0))) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) x) (l_e_st_eq_landau_n_rt_urt ksi x)) y0 (l_e_st_eq_landau_n_rt_rp_5158_t13 ksi x0 ux n y0 o))))))). Time Defined. (* constant 2845 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_not (l_imp (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 x))) (l_e_st_eq_landau_n_rt_rp_5158_t7 ksi x0 ux n) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_not (l_imp (l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux)) (l_e_st_eq_landau_n_rt_lessis x0 x))) => l_e_st_eq_landau_n_rt_rp_5158_t14 ksi x0 ux n x t)))))). Time Defined. (* constant 2846 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => l_e_st_eq_landau_n_rt_rp_moreisi1 (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi (l_e_st_eq_landau_n_rt_rp_5158_t15 ksi x0 ux n))))). Time Defined. (* constant 2847 *) Definition l_e_st_eq_landau_n_rt_rp_satz158b : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => l_imp_th1 (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi) (fun (t:l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0) => l_e_st_eq_landau_n_rt_rp_5158_t4 ksi x0 ux t) (fun (t:l_not (l_e_st_eq_landau_n_rt_min (l_e_st_eq_landau_n_rt_rp_5158_s1 ksi x0 ux) x0)) => l_e_st_eq_landau_n_rt_rp_5158_t16 ksi x0 ux t)))). Time Defined. (* constant 2848 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi), l_not (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi) => l_e_st_eq_landau_n_rt_rp_satz123h (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi l))). Time Defined. (* constant 2849 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi), l_not (l_e_st_eq_landau_n_rt_urt ksi x0)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi) => l_imp_th3 (l_e_st_eq_landau_n_rt_urt ksi x0) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi) (l_e_st_eq_landau_n_rt_rp_5158_t17 ksi x0 l) (fun (t:l_e_st_eq_landau_n_rt_urt ksi x0) => l_e_st_eq_landau_n_rt_rp_satz158b ksi x0 t)))). Time Defined. (* constant 2850 *) Definition l_e_st_eq_landau_n_rt_rp_satz158c : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi), l_e_st_eq_landau_n_rt_lrt ksi x0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi) => l_et (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_rp_5158_t18 ksi x0 l)))). Time Defined. (* constant 2851 *) Definition l_e_st_eq_landau_n_rt_rp_5158_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi), l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi) => l_e_st_eq_landau_n_rt_rp_satz123c (l_e_st_eq_landau_n_rt_rp_5158_xr ksi x0) ksi m))). Time Defined. (* constant 2852 *) Definition l_e_st_eq_landau_n_rt_rp_satz158d : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi), l_e_st_eq_landau_n_rt_urt ksi x0))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi) => l_imp_th3 (l_e_st_eq_landau_n_rt_lrt ksi x0) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) ksi) (l_e_st_eq_landau_n_rt_rp_5158_t19 ksi x0 m) (fun (t:l_e_st_eq_landau_n_rt_lrt ksi x0) => l_e_st_eq_landau_n_rt_rp_satz158a ksi x0 t)))). Time Defined. (* constant 2853 *) Definition l_e_st_eq_landau_n_rt_rp_5159_xr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x0)))). Time Defined. (* constant 2854 *) Definition l_e_st_eq_landau_n_rt_rp_5159_zr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt z0))))))). Time Defined. (* constant 2855 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt eta z0), (forall (k:l_e_st_eq_landau_n_rt_less x0 z0), l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_zr ksi eta l x0 ux lx z0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt eta z0) => (fun (k:l_e_st_eq_landau_n_rt_less x0 z0) => l_e_st_eq_landau_n_rt_rp_satz127a ksi (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) (l_e_st_eq_landau_n_rt_rp_5159_zr ksi eta l x0 ux lx z0) (l_e_st_eq_landau_n_rt_rp_satz124 (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) ksi (l_e_st_eq_landau_n_rt_rp_satz158b ksi x0 ux)) (l_e_st_eq_landau_n_rt_rp_satz154c x0 z0 k)))))))))). Time Defined. (* constant 2856 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt eta z0), (forall (k:l_e_st_eq_landau_n_rt_less x0 z0), l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_zr ksi eta l x0 ux lx z0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_zr ksi eta l x0 ux lx z0) eta)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt eta z0) => (fun (k:l_e_st_eq_landau_n_rt_less x0 z0) => l_andi (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_zr ksi eta l x0 ux lx z0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_zr ksi eta l x0 ux lx z0) eta) (l_e_st_eq_landau_n_rt_rp_5159_t1 ksi eta l x0 ux lx z0 lz k) (l_e_st_eq_landau_n_rt_rp_satz158a eta z0 lz)))))))))). Time Defined. (* constant 2857 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt eta z0), (forall (k:l_e_st_eq_landau_n_rt_less x0 z0), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt eta z0) => (fun (k:l_e_st_eq_landau_n_rt_less x0 z0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta)) z0 (l_e_st_eq_landau_n_rt_rp_5159_t2 ksi eta l x0 ux lx z0 lz k)))))))))). Time Defined. (* constant 2858 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt ksi x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt eta x0), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt ksi x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt eta x0) => l_e_st_eq_landau_n_rt_cutapp3 eta x0 lx (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt eta x) => (fun (u:l_e_st_eq_landau_n_rt_less x0 x) => l_e_st_eq_landau_n_rt_rp_5159_t3 ksi eta l x0 ux lx x t u))))))))). Time Defined. (* constant 2859 *) Definition l_e_st_eq_landau_n_rt_rp_satz159 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_lessapp ksi eta l (l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt ksi x) => (fun (u:l_e_st_eq_landau_n_rt_lrt eta x) => l_e_st_eq_landau_n_rt_rp_5159_t4 ksi eta l x t u)))))). Time Defined. (* constant 2860 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) eta)), l_and3 (l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0)) (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) eta)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) eta)) => l_andi (l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0)) (l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) eta)) (l_e_st_eq_landau_n_rt_rp_ratrpi x0) a))))). Time Defined. (* constant 2861 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) eta)), l_e_st_eq_landau_n_rt_rp_some (fun (c:l_e_st_eq_landau_n_rt_cut) => l_and3 (l_e_st_eq_landau_n_rt_rp_ratrp c) (l_e_st_eq_landau_n_rt_rp_less ksi c) (l_e_st_eq_landau_n_rt_rp_less c eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) eta)) => l_somei l_e_st_eq_landau_n_rt_cut (fun (c:l_e_st_eq_landau_n_rt_cut) => l_and3 (l_e_st_eq_landau_n_rt_rp_ratrp c) (l_e_st_eq_landau_n_rt_rp_less ksi c) (l_e_st_eq_landau_n_rt_rp_less c eta)) (l_e_st_eq_landau_n_rt_rp_5159_xr ksi eta l x0) (l_e_st_eq_landau_n_rt_rp_5159_t5 ksi eta l x0 a)))))). Time Defined. (* constant 2862 *) Definition l_e_st_eq_landau_n_rt_rp_satz159a : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_some (fun (a:l_e_st_eq_landau_n_rt_cut) => l_and3 (l_e_st_eq_landau_n_rt_rp_ratrp a) (l_e_st_eq_landau_n_rt_rp_less ksi a) (l_e_st_eq_landau_n_rt_rp_less a eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta)) (l_e_st_eq_landau_n_rt_rp_satz159 ksi eta l) (l_e_st_eq_landau_n_rt_rp_some (fun (a:l_e_st_eq_landau_n_rt_cut) => l_and3 (l_e_st_eq_landau_n_rt_rp_ratrp a) (l_e_st_eq_landau_n_rt_rp_less ksi a) (l_e_st_eq_landau_n_rt_rp_less a eta))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta)) => l_e_st_eq_landau_n_rt_rp_5159_t6 ksi eta l x t))))). Time Defined. (* constant 2863 *) Definition l_e_st_eq_landau_n_rt_rp_5159_yr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt y0)))))). Time Defined. (* constant 2864 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta)), l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta) a))))))). Time Defined. (* constant 2865 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta)), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta) a))))))). Time Defined. (* constant 2866 *) Definition l_e_st_eq_landau_n_rt_rp_5159_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta)), p))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5159_yr ksi eta l p p1 y0) eta)) => p1 y0 (l_e_st_eq_landau_n_rt_rp_5159_t7 ksi eta l p p1 y0 a) (l_e_st_eq_landau_n_rt_rp_5159_t8 ksi eta l p p1 y0 a)))))))). Time Defined. (* constant 2867 *) Definition l_e_st_eq_landau_n_rt_rp_satz159app : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))), p))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)), (forall (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta), p)))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta)) (l_e_st_eq_landau_n_rt_rp_satz159 ksi eta l) p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) eta)) => l_e_st_eq_landau_n_rt_rp_5159_t9 ksi eta l p p1 x t))))))). Time Defined. (* constant 2868 *) Definition l_e_st_eq_landau_n_rt_rp_5160_zr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => l_e_st_eq_landau_n_rt_rp_rpofrt z0)))). Time Defined. (* constant 2869 *) Definition l_e_st_eq_landau_n_rt_rp_5160_nm : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_5160_zr ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_ts ksi eta) m)))). Time Defined. (* constant 2870 *) Definition l_e_st_eq_landau_n_rt_rp_5160_dn : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) l_e_st_eq_landau_n_rt_rp_1rp)))). Time Defined. (* constant 2871 *) Definition l_e_st_eq_landau_n_rt_rp_5160_fr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => l_e_st_eq_landau_n_rt_rp_ov (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m))))). Time Defined. (* constant 2872 *) Definition l_e_st_eq_landau_n_rt_rp_5160_zeta : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => l_e_ite (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)))). Time Defined. (* constant 2873 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) => l_e_itet (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp l))))). Time Defined. (* constant 2874 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) => l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_t1 ksi eta z0 m l)))))). Time Defined. (* constant 2875 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) => l_e_st_eq_landau_n_rt_rp_lessisi1 (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_satz127a (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5160_t2 ksi eta z0 m l) l)))))). Time Defined. (* constant 2876 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)) => l_e_itef (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp n))))). Time Defined. (* constant 2877 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)) => l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5160_t4 ksi eta z0 m n)))))). Time Defined. (* constant 2878 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)) => l_e_st_eq_landau_n_rt_rp_trlessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_t5 ksi eta z0 m n) (l_e_st_eq_landau_n_rt_rp_satz124 (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_satz123f (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp n))))))). Time Defined. (* constant 2879 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) (fun (t:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) => l_e_st_eq_landau_n_rt_rp_5160_t3 ksi eta z0 m t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)) => l_e_st_eq_landau_n_rt_rp_5160_t5 ksi eta z0 m t))))). Time Defined. (* constant 2880 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m)) (fun (t:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp) => l_e_st_eq_landau_n_rt_rp_5160_t2 ksi eta z0 m t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp)) => l_e_st_eq_landau_n_rt_rp_5160_t6 ksi eta z0 m t))))). Time Defined. (* constant 2881 *) Definition l_e_st_eq_landau_n_rt_rp_5160_zr1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt z1))))). Time Defined. (* constant 2882 *) Definition l_e_st_eq_landau_n_rt_rp_5160_zr2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt z2)))))))). Time Defined. (* constant 2883 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_disttp2 (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_satz147a (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) l2 l4))))))))))). Time Defined. (* constant 2884 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) eta)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_disttp1 ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) eta) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) eta))))))))))))). Time Defined. (* constant 2885 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_satz149a (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_satz139a ksi ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_lessisi2 ksi ksi (l_e_refis l_e_st_eq_landau_n_rt_cut ksi)) (l_e_st_eq_landau_n_rt_rp_5160_t7 ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))))))))))))). Time Defined. (* constant 2886 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_satz139a (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_t10 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_t11 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2887 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_satz127b (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_5160_t9 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_t12 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2888 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl eta ksi) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl ksi eta) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_asspl2 eta ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl eta ksi) (l_e_st_eq_landau_n_rt_rp_pl ksi eta) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_compl eta ksi)))))))))))). Time Defined. (* constant 2889 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_distpt1 eta (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ists1 (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_t14 ksi eta z0 m z1 l1 l2 z2 l3 l4)))))))))))). Time Defined. (* constant 2890 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_5160_t15 ksi eta z0 m z1 l1 l2 z2 l3 l4)))))))))))). Time Defined. (* constant 2891 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl ksi l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_5160_t16 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_t13 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2892 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_islessis12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_satz153e (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_satz149a (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_fr ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_t8 ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m)))))))))))))). Time Defined. (* constant 2893 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t19 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_satz139a (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts ksi eta))) (l_e_st_eq_landau_n_rt_rp_5160_t18 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2894 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m)))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_satz127b (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_dn ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_t17 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_t19 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2895 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_5160_zr ksi eta z0 m))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_5160_nm ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_zr ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_satz140c (l_e_st_eq_landau_n_rt_rp_5160_zr ksi eta z0 m) (l_e_st_eq_landau_n_rt_rp_ts ksi eta) m) (l_e_st_eq_landau_n_rt_rp_5160_t20 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2896 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_less (l_e_st_eq_landau_n_rt_ts z1 z2) z0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_satz154f (l_e_st_eq_landau_n_rt_ts z1 z2) z0 (l_e_st_eq_landau_n_rt_rp_isless1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts z1 z2)) (l_e_st_eq_landau_n_rt_rp_5160_zr ksi eta z0 m) (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts z1 z2)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) (l_e_st_eq_landau_n_rt_rp_satz155c z1 z2)) (l_e_st_eq_landau_n_rt_rp_5160_t21 ksi eta z0 m z1 l1 l2 z2 l3 l4)))))))))))). Time Defined. (* constant 2897 *) Definition l_e_st_eq_landau_n_rt_rp_5160_x0 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rat)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_ov z0 z2)))))))))). Time Defined. (* constant 2898 *) Definition l_e_st_eq_landau_n_rt_rp_5160_xr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_cut)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2899 *) Definition l_e_st_eq_landau_n_rt_rp_5160_y0 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rat)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => z2)))))))))). Time Defined. (* constant 2900 *) Definition l_e_st_eq_landau_n_rt_rp_5160_yr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_cut)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5160_y0 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2901 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_y0 ksi eta z0 m z1 l1 l2 z2 l3 l4)) z0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_satz110e z0 z2)))))))))). Time Defined. (* constant 2902 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t24 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) z2) (l_e_st_eq_landau_n_rt_ts z1 z2))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_ismore1 z0 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) z2) (l_e_st_eq_landau_n_rt_ts z1 z2) (l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) z2) z0 (l_e_st_eq_landau_n_rt_rp_5160_t23 ksi eta z0 m z1 l1 l2 z2 l3 l4)) (l_e_st_eq_landau_n_rt_satz83 (l_e_st_eq_landau_n_rt_ts z1 z2) z0 (l_e_st_eq_landau_n_rt_rp_5160_t22 ksi eta z0 m z1 l1 l2 z2 l3 l4)))))))))))). Time Defined. (* constant 2903 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t25 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) z1)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_satz106a (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) z1 z2 (l_e_st_eq_landau_n_rt_rp_5160_t24 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2904 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t26 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_xr ksi eta z0 m z1 l1 l2 z2 l3 l4) ksi)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_trmore (l_e_st_eq_landau_n_rt_rp_5160_xr ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) ksi (l_e_st_eq_landau_n_rt_rp_satz154a (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) z1 (l_e_st_eq_landau_n_rt_rp_5160_t25 ksi eta z0 m z1 l1 l2 z2 l3 l4)) (l_e_st_eq_landau_n_rt_rp_satz122 ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) l1))))))))))). Time Defined. (* constant 2905 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t27 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_yr ksi eta z0 m z1 l1 l2 z2 l3 l4) eta)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_satz122 eta (l_e_st_eq_landau_n_rt_rp_5160_yr ksi eta z0 m z1 l1 l2 z2 l3 l4) l3)))))))))). Time Defined. (* constant 2906 *) Definition l_e_st_eq_landau_n_rt_rp_5160_ur : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt u0)))). Time Defined. (* constant 2907 *) Definition l_e_st_eq_landau_n_rt_rp_5160_vr : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (v0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt v0))))). Time Defined. (* constant 2908 *) Definition l_e_st_eq_landau_n_rt_rp_5160_prop1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (v0:l_e_st_eq_landau_n_rt_rat), Prop))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (v0:l_e_st_eq_landau_n_rt_rat) => l_and3 (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_ur ksi eta z0 u0) ksi) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_vr ksi eta z0 u0 v0) eta) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts u0 v0) z0)))))). Time Defined. (* constant 2909 *) Definition l_e_st_eq_landau_n_rt_rp_5160_prop2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), Prop))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x y))))). Time Defined. (* constant 2910 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t28 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_y0 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_and3i (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_xr ksi eta z0 m z1 l1 l2 z2 l3 l4) ksi) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_yr ksi eta z0 m z1 l1 l2 z2 l3 l4) eta) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_y0 ksi eta z0 m z1 l1 l2 z2 l3 l4)) z0) (l_e_st_eq_landau_n_rt_rp_5160_t26 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_t27 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_t23 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2911 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t29 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) y))))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) y) (l_e_st_eq_landau_n_rt_rp_5160_y0 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_t28 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2912 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t30 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_5160_prop2 ksi eta z0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr2 ksi eta z0 m z1 l1 l2 z2) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x y)) (l_e_st_eq_landau_n_rt_rp_5160_x0 ksi eta z0 m z1 l1 l2 z2 l3 l4) (l_e_st_eq_landau_n_rt_rp_5160_t29 ksi eta z0 m z1 l1 l2 z2 l3 l4))))))))))). Time Defined. (* constant 2913 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t31 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))), l_e_st_eq_landau_n_rt_rp_5160_prop2 ksi eta z0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5160_zr1 ksi eta z0 m z1) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_satz159app eta (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_satz133a eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_prop2 ksi eta z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_less eta (l_e_st_eq_landau_n_rt_rp_rpofrt x)) => (fun (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_pl eta (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_5160_t30 ksi eta z0 m z1 l1 l2 x t u)))))))))). Time Defined. (* constant 2914 *) Definition l_e_st_eq_landau_n_rt_rp_satz160 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), l_e_st_eq_landau_n_rt_some (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and3 (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => l_e_st_eq_landau_n_rt_rp_satz159app ksi (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_satz133a ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m)) (l_e_st_eq_landau_n_rt_rp_5160_prop2 ksi eta z0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_less ksi (l_e_st_eq_landau_n_rt_rp_rpofrt x)) => (fun (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_pl ksi (l_e_st_eq_landau_n_rt_rp_5160_zeta ksi eta z0 m))) => l_e_st_eq_landau_n_rt_rp_5160_t31 ksi eta z0 m x t u))))))). Time Defined. (* constant 2915 *) Definition l_e_st_eq_landau_n_rt_rp_5160_xr1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))), (forall (x1:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x1))))))). Time Defined. (* constant 2916 *) Definition l_e_st_eq_landau_n_rt_rp_5160_yr1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)), (forall (y1:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt y1))))))))). Time Defined. (* constant 2917 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t32 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y1), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_xr1 ksi eta z0 m p p1 x1) ksi)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y1) => l_and3e1 (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_xr1 ksi eta z0 m p p1 x1) ksi) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_yr1 ksi eta z0 m p p1 x1 px y1) eta) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x1 y1) z0) py)))))))))). Time Defined. (* constant 2918 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t33 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y1), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_yr1 ksi eta z0 m p p1 x1 px y1) eta)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y1) => l_and3e2 (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_xr1 ksi eta z0 m p p1 x1) ksi) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_yr1 ksi eta z0 m p p1 x1 px y1) eta) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x1 y1) z0) py)))))))))). Time Defined. (* constant 2919 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t34 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y1), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x1 y1) z0)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y1) => l_and3e3 (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_xr1 ksi eta z0 m p p1 x1) ksi) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5160_yr1 ksi eta z0 m p p1 x1 px y1) eta) (l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x1 y1) z0) py)))))))))). Time Defined. (* constant 2920 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t35 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (py:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y1), p)))))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (py:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y1) => p1 x1 (l_e_st_eq_landau_n_rt_rp_5160_t32 ksi eta z0 m p p1 x1 px y1 py) y1 (l_e_st_eq_landau_n_rt_rp_5160_t33 ksi eta z0 m p p1 x1 px y1 py) (l_e_st_eq_landau_n_rt_rp_5160_t34 ksi eta z0 m p p1 x1 px y1 py))))))))))). Time Defined. (* constant 2921 *) Definition l_e_st_eq_landau_n_rt_rp_5160_t36 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)), p)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (px:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y)) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y) px p (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (v:l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x1 y) => l_e_st_eq_landau_n_rt_rp_5160_t35 ksi eta z0 m p p1 x1 px y v)))))))))). Time Defined. (* constant 2922 *) Definition l_e_st_eq_landau_n_rt_rp_satz160app : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))), p)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt z0) (l_e_st_eq_landau_n_rt_rp_ts ksi eta)) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rat), (forall (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) ksi), (forall (y:l_e_st_eq_landau_n_rt_rat), (forall (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) eta), (forall (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z0), p)))))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x y)) (l_e_st_eq_landau_n_rt_rp_satz160 ksi eta z0 m) p (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5160_prop1 ksi eta z0 x y)) => l_e_st_eq_landau_n_rt_rp_5160_t36 ksi eta z0 m p p1 x t)))))))). Time Defined. (* constant 2923 *) Definition l_e_st_eq_landau_n_rt_rp_5161_min : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_ite (l_e_st_eq_landau_n_rt_rp_less ksi eta) l_e_st_eq_landau_n_rt_cut ksi eta)). Time Defined. (* constant 2924 *) Definition l_e_st_eq_landau_n_rt_rp_5161_max : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => l_e_ite (l_e_st_eq_landau_n_rt_rp_more ksi eta) l_e_st_eq_landau_n_rt_cut ksi eta)). Time Defined. (* constant 2925 *) Definition l_e_st_eq_landau_n_rt_rp_5161_ur : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt u0))). Time Defined. (* constant 2926 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0) => l_e_st_eq_landau_n_rt_rp_satz158a (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0 lu)))). Time Defined. (* constant 2927 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) ksi (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) (l_e_itet (l_e_st_eq_landau_n_rt_rp_less ksi eta) l_e_st_eq_landau_n_rt_cut ksi eta l) (l_e_st_eq_landau_n_rt_rp_5161_t1 ksi eta u0 lu)))))). Time Defined. (* constant 2928 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0), (forall (l:l_e_st_eq_landau_n_rt_rp_less ksi eta), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0) => (fun (l:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_trless (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi eta (l_e_st_eq_landau_n_rt_rp_5161_t2 ksi eta u0 lu l) l))))). Time Defined. (* constant 2929 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) eta (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) (l_e_itef (l_e_st_eq_landau_n_rt_rp_less ksi eta) l_e_st_eq_landau_n_rt_cut ksi eta n) (l_e_st_eq_landau_n_rt_rp_5161_t1 ksi eta u0 lu)))))). Time Defined. (* constant 2930 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)) => l_e_st_eq_landau_n_rt_rp_satz127b (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta ksi (l_e_st_eq_landau_n_rt_rp_5161_t4 ksi eta u0 lu n) (l_e_st_eq_landau_n_rt_rp_satz124 ksi eta (l_e_st_eq_landau_n_rt_rp_satz123f ksi eta n))))))). Time Defined. (* constant 2931 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi) (fun (t:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_5161_t2 ksi eta u0 lu t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)) => l_e_st_eq_landau_n_rt_rp_5161_t5 ksi eta u0 lu t))))). Time Defined. (* constant 2932 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (lu:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min ksi eta) u0) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_less ksi eta) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta) (fun (t:l_e_st_eq_landau_n_rt_rp_less ksi eta) => l_e_st_eq_landau_n_rt_rp_5161_t3 ksi eta u0 lu t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_less ksi eta)) => l_e_st_eq_landau_n_rt_rp_5161_t4 ksi eta u0 lu t))))). Time Defined. (* constant 2933 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0) => l_e_st_eq_landau_n_rt_rp_satz158b (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0 uu)))). Time Defined. (* constant 2934 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_ismoreis2 (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) ksi (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) (l_e_itet (l_e_st_eq_landau_n_rt_rp_more ksi eta) l_e_st_eq_landau_n_rt_cut ksi eta m) (l_e_st_eq_landau_n_rt_rp_5161_t8 ksi eta u0 uu)))))). Time Defined. (* constant 2935 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_trmoreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi eta (l_e_st_eq_landau_n_rt_rp_5161_t9 ksi eta u0 uu m) (l_e_st_eq_landau_n_rt_rp_moreisi1 ksi eta m)))))). Time Defined. (* constant 2936 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)) => l_e_st_eq_landau_n_rt_rp_ismoreis2 (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) eta (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) (l_e_itef (l_e_st_eq_landau_n_rt_rp_more ksi eta) l_e_st_eq_landau_n_rt_cut ksi eta n) (l_e_st_eq_landau_n_rt_rp_5161_t8 ksi eta u0 uu)))))). Time Defined. (* constant 2937 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)) => l_e_st_eq_landau_n_rt_rp_trmoreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta ksi (l_e_st_eq_landau_n_rt_rp_5161_t11 ksi eta u0 uu n) (l_e_st_eq_landau_n_rt_rp_satz125 ksi eta (l_e_st_eq_landau_n_rt_rp_satz123e ksi eta n))))))). Time Defined. (* constant 2938 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) ksi) (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_5161_t9 ksi eta u0 uu t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)) => l_e_st_eq_landau_n_rt_rp_5161_t12 ksi eta u0 uu t))))). Time Defined. (* constant 2939 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (u0:l_e_st_eq_landau_n_rt_rat), (forall (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (u0:l_e_st_eq_landau_n_rt_rat) => (fun (uu:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max ksi eta) u0) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_more ksi eta) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_ur ksi eta u0) eta) (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_5161_t10 ksi eta u0 uu t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_more ksi eta)) => l_e_st_eq_landau_n_rt_rp_5161_t11 ksi eta u0 uu t))))). Time Defined. (* constant 2940 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t15 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi1:l_e_st_eq_landau_n_rt_cut), (forall (ksi2:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi1 ksi2), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts ksi1 ksi1) (l_e_st_eq_landau_n_rt_rp_ts ksi2 ksi2))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi1:l_e_st_eq_landau_n_rt_cut) => (fun (ksi2:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi1 ksi2) => l_e_st_eq_landau_n_rt_rp_satz147 ksi1 ksi2 ksi1 ksi2 m m)))). Time Defined. (* constant 2941 *) Definition l_e_st_eq_landau_n_rt_rp_5161_sq1 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi1:l_e_st_eq_landau_n_rt_cut), (forall (ksi2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi1:l_e_st_eq_landau_n_rt_cut) => (fun (ksi2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ts ksi1 ksi1))). Time Defined. (* constant 2942 *) Definition l_e_st_eq_landau_n_rt_rp_5161_sq2 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi1:l_e_st_eq_landau_n_rt_cut), (forall (ksi2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi1:l_e_st_eq_landau_n_rt_cut) => (fun (ksi2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ts ksi2 ksi2))). Time Defined. (* constant 2943 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t16 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi1:l_e_st_eq_landau_n_rt_cut), (forall (ksi2:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi1 ksi2), l_e_st_eq_landau_n_rt_rp_nis (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi1:l_e_st_eq_landau_n_rt_cut) => (fun (ksi2:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi1 ksi2) => l_ec3e21 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2)) (l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2)) (l_e_st_eq_landau_n_rt_rp_5161_t15 zeta ksi1 ksi2 m))))). Time Defined. (* constant 2944 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t17 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi1:l_e_st_eq_landau_n_rt_cut), (forall (ksi2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) zeta), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi1:l_e_st_eq_landau_n_rt_cut) => (fun (ksi2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) zeta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta) => l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta i j))))). Time Defined. (* constant 2945 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t18 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi1:l_e_st_eq_landau_n_rt_cut), (forall (ksi2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) zeta), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta), l_not (l_e_st_eq_landau_n_rt_rp_more ksi1 ksi2)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi1:l_e_st_eq_landau_n_rt_cut) => (fun (ksi2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) zeta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta) => (fun (t:l_e_st_eq_landau_n_rt_rp_more ksi1 ksi2) => l_e_st_eq_landau_n_rt_rp_5161_t16 zeta ksi1 ksi2 t (l_e_st_eq_landau_n_rt_rp_5161_t17 zeta ksi1 ksi2 i j))))))). Time Defined. (* constant 2946 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t19 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi1:l_e_st_eq_landau_n_rt_cut), (forall (ksi2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) zeta), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta), l_not (l_e_st_eq_landau_n_rt_rp_less ksi1 ksi2)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi1:l_e_st_eq_landau_n_rt_cut) => (fun (ksi2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) zeta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta) => (fun (t:l_e_st_eq_landau_n_rt_rp_less ksi1 ksi2) => l_e_st_eq_landau_n_rt_rp_5161_t16 zeta ksi2 ksi1 (l_e_st_eq_landau_n_rt_rp_satz122 ksi1 ksi2 t) (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_t17 zeta ksi1 ksi2 i j)))))))). Time Defined. (* constant 2947 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t20 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi1:l_e_st_eq_landau_n_rt_cut), (forall (ksi2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) zeta), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta), l_e_st_eq_landau_n_rt_rp_is ksi1 ksi2))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi1:l_e_st_eq_landau_n_rt_cut) => (fun (ksi2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq1 zeta ksi1 ksi2) zeta) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_sq2 zeta ksi1 ksi2) zeta) => l_or3e1 (l_e_st_eq_landau_n_rt_rp_is ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_more ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_less ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_satz123a ksi1 ksi2) (l_e_st_eq_landau_n_rt_rp_5161_t18 zeta ksi1 ksi2 i j) (l_e_st_eq_landau_n_rt_rp_5161_t19 zeta ksi1 ksi2 i j)))))). Time Defined. (* constant 2948 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t21 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_amone l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) zeta)). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_e_st_eq_landau_n_rt_cut) => (fun (b:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) zeta) => (fun (u:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts b b) zeta) => l_e_st_eq_landau_n_rt_rp_5161_t20 zeta a b t u))))). Time Defined. (* constant 2949 *) Definition l_e_st_eq_landau_n_rt_rp_5161_sqrtset : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_set l_e_st_eq_landau_n_rt_rat). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) zeta)). Time Defined. (* constant 2950 *) Definition l_e_st_eq_landau_n_rt_rp_5161_xr : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x0)). Time Defined. (* constant 2951 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t22 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_5161_t6 l_e_st_eq_landau_n_rt_rp_1rp zeta x0 lx))). Time Defined. (* constant 2952 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t23 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) zeta))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_5161_t7 l_e_st_eq_landau_n_rt_rp_1rp zeta x0 lx))). Time Defined. (* constant 2953 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t24 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) zeta))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_isless1 (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) zeta (l_e_st_eq_landau_n_rt_rp_satz151a (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_5161_t23 zeta x0 lx)))). Time Defined. (* constant 2954 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t25 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) zeta))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_trless (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) zeta (l_e_st_eq_landau_n_rt_rp_satz148c (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0))) (l_e_st_eq_landau_n_rt_rp_5161_t22 zeta x0 lx)) (l_e_st_eq_landau_n_rt_rp_5161_t24 zeta x0 lx)))). Time Defined. (* constant 2955 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t26 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) zeta) x0 (l_e_st_eq_landau_n_rt_rp_5161_t25 zeta x0 lx)))). Time Defined. (* constant 2956 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t27 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_5161_t13 l_e_st_eq_landau_n_rt_rp_1rp zeta x0 ux))). Time Defined. (* constant 2957 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t28 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) zeta))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_5161_t14 l_e_st_eq_landau_n_rt_rp_1rp zeta x0 ux))). Time Defined. (* constant 2958 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t29 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) zeta))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_ismoreis1 (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) zeta (l_e_st_eq_landau_n_rt_rp_satz151a (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_5161_t28 zeta x0 ux)))). Time Defined. (* constant 2959 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t30 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) zeta))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_trmoreis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) zeta (l_e_st_eq_landau_n_rt_rp_satz149 (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_moreisi2 (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0))) (l_e_st_eq_landau_n_rt_rp_5161_t27 zeta x0 ux)) (l_e_st_eq_landau_n_rt_rp_5161_t29 zeta x0 ux)))). Time Defined. (* constant 2960 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t31 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) zeta)))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_rp_satz123c (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) zeta (l_e_st_eq_landau_n_rt_rp_5161_t30 zeta x0 ux)))). Time Defined. (* constant 2961 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t32 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_not (l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_imp_th3 (l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) zeta) (l_e_st_eq_landau_n_rt_rp_5161_t31 zeta x0 ux) (fun (t:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) zeta) x0 t)))). Time Defined. (* constant 2962 *) Definition l_e_st_eq_landau_n_rt_rp_5161_yr : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (y0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). Time Defined. (* constant 2963 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t33 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) zeta))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) zeta) x0 i))). Time Defined. (* constant 2964 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t34 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_yr zeta x0 i y0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_rp_satz154c y0 x0 l))))). Time Defined. (* constant 2965 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t35 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yr zeta x0 i y0) (l_e_st_eq_landau_n_rt_rp_5161_yr zeta x0 i y0)) zeta))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_eq_landau_n_rt_rp_trless (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yr zeta x0 i y0) (l_e_st_eq_landau_n_rt_rp_5161_yr zeta x0 i y0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) zeta (l_e_st_eq_landau_n_rt_rp_satz147a (l_e_st_eq_landau_n_rt_rp_5161_yr zeta x0 i y0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_yr zeta x0 i y0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_t34 zeta x0 i y0 l) (l_e_st_eq_landau_n_rt_rp_5161_t34 zeta x0 i y0 l)) (l_e_st_eq_landau_n_rt_rp_5161_t33 zeta x0 i)))))). Time Defined. (* constant 2966 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t36 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) zeta) y0 (l_e_st_eq_landau_n_rt_rp_5161_t35 zeta x0 i y0 l)))))). Time Defined. (* constant 2967 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t37 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), l_e_st_eq_landau_n_rt_rp_more zeta (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) zeta (l_e_st_eq_landau_n_rt_rp_5161_t33 zeta x0 i)))). Time Defined. (* constant 2968 *) Definition l_e_st_eq_landau_n_rt_rp_5161_nm : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), l_e_st_eq_landau_n_rt_cut))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_eq_landau_n_rt_rp_mn zeta (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_5161_t37 zeta x0 i)))). Time Defined. (* constant 2969 *) Definition l_e_st_eq_landau_n_rt_rp_5161_dn : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), l_e_st_eq_landau_n_rt_cut))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp)))). Time Defined. (* constant 2970 *) Definition l_e_st_eq_landau_n_rt_rp_5161_fr : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), l_e_st_eq_landau_n_rt_cut))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_eq_landau_n_rt_rp_ov (l_e_st_eq_landau_n_rt_rp_5161_nm zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i)))). Time Defined. (* constant 2971 *) Definition l_e_st_eq_landau_n_rt_rp_5161_zr : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt z0)))). Time Defined. (* constant 2972 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t38 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) l_e_st_eq_landau_n_rt_rp_1rp))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_5161_t6 l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i) z0 lz))))). Time Defined. (* constant 2973 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t39 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_5161_t7 l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i) z0 lz))))). Time Defined. (* constant 2974 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t40 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) x0))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_satz94 x0 z0))))). Time Defined. (* constant 2975 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t41 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ists12 (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_satz155a x0 z0) (l_e_st_eq_landau_n_rt_rp_satz155a x0 z0)) (l_e_st_eq_landau_n_rt_rp_disttp2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))))))). Time Defined. (* constant 2976 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t42 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_5161_t41 zeta x0 i z0 lz)))))). Time Defined. (* constant 2977 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t43 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_disttp1 (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0))))))))). Time Defined. (* constant 2978 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t44 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_satz145c (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) (l_e_st_eq_landau_n_rt_rp_satz138c (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0))) (l_e_st_eq_landau_n_rt_rp_5161_t38 zeta x0 i z0 lz))))))). Time Defined. (* constant 2979 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t45 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_satz138c (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_t43 zeta x0 i z0 lz) (l_e_st_eq_landau_n_rt_rp_5161_t44 zeta x0 i z0 lz)))))). Time Defined. (* constant 2980 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t46 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_distpt1 (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)))))))). Time Defined. (* constant 2981 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t47 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_5161_t42 zeta x0 i z0 lz) (l_e_st_eq_landau_n_rt_rp_5161_t46 zeta x0 i z0 lz) (l_e_st_eq_landau_n_rt_rp_5161_t45 zeta x0 i z0 lz)))))). Time Defined. (* constant 2982 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t48 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_nm zeta x0 i)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) (l_e_st_eq_landau_n_rt_rp_5161_nm zeta x0 i) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_satz153c (l_e_st_eq_landau_n_rt_rp_5161_nm zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i)) (l_e_st_eq_landau_n_rt_rp_satz148c (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0) (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i) (l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i))) (l_e_st_eq_landau_n_rt_rp_5161_t39 zeta x0 i z0 lz))))))). Time Defined. (* constant 2983 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t49 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_5161_nm zeta x0 i))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_satz138c (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0)) (l_e_st_eq_landau_n_rt_rp_5161_nm zeta x0 i) (l_e_st_eq_landau_n_rt_rp_lessisi2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)))) (l_e_st_eq_landau_n_rt_rp_5161_t48 zeta x0 i z0 lz)))))). Time Defined. (* constant 2984 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t50 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) zeta))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_5161_nm zeta x0 i)) zeta (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) (l_e_st_eq_landau_n_rt_rp_satz140c zeta (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_5161_t37 zeta x0 i)) (l_e_st_eq_landau_n_rt_rp_5161_t49 zeta x0 i z0 lz)))))). Time Defined. (* constant 2985 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t51 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0))) zeta))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_eq_landau_n_rt_rp_trless (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl x0 z0))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0) (l_e_st_eq_landau_n_rt_rp_5161_xr zeta x0)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_dn zeta x0 i) (l_e_st_eq_landau_n_rt_rp_5161_zr zeta x0 i z0))) zeta (l_e_st_eq_landau_n_rt_rp_5161_t47 zeta x0 i z0 lz) (l_e_st_eq_landau_n_rt_rp_5161_t50 zeta x0 i z0 lz)))))). Time Defined. (* constant 2986 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t52 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) zeta) (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_rp_5161_t51 zeta x0 i z0 lz)))))). Time Defined. (* constant 2987 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t53 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_and (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) x0)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_andi (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_pl x0 z0) x0) (l_e_st_eq_landau_n_rt_rp_5161_t52 zeta x0 i z0 lz) (l_e_st_eq_landau_n_rt_rp_5161_t40 zeta x0 i z0 lz)))))). Time Defined. (* constant 2988 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t54 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), (forall (z0:l_e_st_eq_landau_n_rt_rat), (forall (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_more y x0))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (z0:l_e_st_eq_landau_n_rt_rat) => (fun (lz:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) z0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_more y x0)) (l_e_st_eq_landau_n_rt_pl x0 z0) (l_e_st_eq_landau_n_rt_rp_5161_t53 zeta x0 i z0 lz)))))). Time Defined. (* constant 2989 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t55 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_more y x0))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_eq_landau_n_rt_cutapp1a (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) (l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_more y x0))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_5161_fr zeta x0 i)) x) => l_e_st_eq_landau_n_rt_rp_5161_t54 zeta x0 i x t))))). Time Defined. (* constant 2990 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t56 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) y0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) y0) => l_e_st_eq_landau_n_rt_cut2 (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta) x0 (l_e_st_eq_landau_n_rt_rp_5161_t26 zeta x0 lx) y0 (l_e_st_eq_landau_n_rt_rp_5161_t32 zeta y0 uy) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_less y x) => l_e_st_eq_landau_n_rt_rp_5161_t36 zeta x t y u)))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_eq_landau_n_rt_rp_5161_t55 zeta x t))))))). Time Defined. (* constant 2991 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t57 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x0) => l_e_st_eq_landau_n_rt_cutapp1b (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_max l_e_st_eq_landau_n_rt_rp_1rp zeta) y) => l_e_st_eq_landau_n_rt_rp_5161_t56 zeta x0 lx y t))))). Time Defined. (* constant 2992 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t58 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_cutapp1a (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_min l_e_st_eq_landau_n_rt_rp_1rp zeta) x) => l_e_st_eq_landau_n_rt_rp_5161_t57 zeta x t))). Time Defined. (* constant 2993 *) Definition l_e_st_eq_landau_n_rt_rp_5161_rtc : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_cutof (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta) (l_e_st_eq_landau_n_rt_rp_5161_t58 zeta)). Time Defined. (* constant 2994 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t59 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_lessis x0 y0), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_lessis x0 y0) => l_or_th9 (l_e_st_eq_landau_n_rt_less x0 y0) (l_e_st_eq_landau_n_rt_is x0 y0) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) l (fun (t:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154c x0 y0 t) (fun (t:l_e_st_eq_landau_n_rt_is x0 y0) => l_e_st_eq_landau_n_rt_rp_satz154b x0 y0 t)))). Time Defined. (* constant 2995 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t60 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (m:l_e_st_eq_landau_n_rt_moreis x0 y0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (m:l_e_st_eq_landau_n_rt_moreis x0 y0) => l_e_st_eq_landau_n_rt_rp_satz125 (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_5161_t59 y0 x0 (l_e_st_eq_landau_n_rt_satz84 x0 y0 m))))). Time Defined. (* constant 2996 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t61 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta m)). Time Defined. (* constant 2997 *) Definition l_e_st_eq_landau_n_rt_rp_5161_zr1 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt z1))). Time Defined. (* constant 2998 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t62 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) z1))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => l_e_st_eq_landau_n_rt_rp_satz158c (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) z1 l2))))). Time Defined. (* constant 2999 *) Definition l_e_st_eq_landau_n_rt_rp_5161_xr1 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x1)))))). Time Defined. (* constant 3000 *) Definition l_e_st_eq_landau_n_rt_rp_5161_xr2 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x2)))))))). Time Defined. (* constant 3001 *) Definition l_e_st_eq_landau_n_rt_rp_5161_xm : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_rat)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_ite (l_e_st_eq_landau_n_rt_more x1 x2) l_e_st_eq_landau_n_rt_rat x1 x2)))))))))). Time Defined. (* constant 3002 *) Definition l_e_st_eq_landau_n_rt_rp_5161_xrm : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_cut)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). Time Defined. (* constant 3003 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t63 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), (forall (o:l_e_st_eq_landau_n_rt_more x1 x2), l_e_st_eq_landau_n_rt_is x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => (fun (o:l_e_st_eq_landau_n_rt_more x1 x2) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) x1 (l_e_itet (l_e_st_eq_landau_n_rt_more x1 x2) l_e_st_eq_landau_n_rt_rat x1 x2 o)))))))))))). Time Defined. (* constant 3004 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t64 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), (forall (o:l_e_st_eq_landau_n_rt_more x1 x2), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => (fun (o:l_e_st_eq_landau_n_rt_more x1 x2) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x) x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) lx1 (l_e_st_eq_landau_n_rt_rp_5161_t63 zeta m z1 l1 l2 x1 lx1 x2 lx2 i o)))))))))))). Time Defined. (* constant 3005 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t65 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), (forall (o:l_e_st_eq_landau_n_rt_more x1 x2), l_e_st_eq_landau_n_rt_lessis x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => (fun (o:l_e_st_eq_landau_n_rt_more x1 x2) => l_e_st_eq_landau_n_rt_lessisi2 x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t63 zeta m z1 l1 l2 x1 lx1 x2 lx2 i o)))))))))))). Time Defined. (* constant 3006 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t66 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), (forall (o:l_e_st_eq_landau_n_rt_more x1 x2), l_e_st_eq_landau_n_rt_lessis x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => (fun (o:l_e_st_eq_landau_n_rt_more x1 x2) => l_e_st_eq_landau_n_rt_lessisi1 x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_satz87b x2 x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_satz82 x1 x2 o) (l_e_st_eq_landau_n_rt_rp_5161_t65 zeta m z1 l1 l2 x1 lx1 x2 lx2 i o))))))))))))). Time Defined. (* constant 3007 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t67 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), (forall (n:l_not (l_e_st_eq_landau_n_rt_more x1 x2)), l_e_st_eq_landau_n_rt_is x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_more x1 x2)) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) x2 (l_e_itef (l_e_st_eq_landau_n_rt_more x1 x2) l_e_st_eq_landau_n_rt_rat x1 x2 n)))))))))))). Time Defined. (* constant 3008 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t68 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), (forall (n:l_not (l_e_st_eq_landau_n_rt_more x1 x2)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_more x1 x2)) => l_e_isp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x) x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) lx2 (l_e_st_eq_landau_n_rt_rp_5161_t67 zeta m z1 l1 l2 x1 lx1 x2 lx2 i n)))))))))))). Time Defined. (* constant 3009 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t69 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), (forall (n:l_not (l_e_st_eq_landau_n_rt_more x1 x2)), l_e_st_eq_landau_n_rt_lessis x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_more x1 x2)) => l_e_st_eq_landau_n_rt_lessisi2 x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t67 zeta m z1 l1 l2 x1 lx1 x2 lx2 i n)))))))))))). Time Defined. (* constant 3010 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t70 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), (forall (n:l_not (l_e_st_eq_landau_n_rt_more x1 x2)), l_e_st_eq_landau_n_rt_lessis x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_more x1 x2)) => l_e_st_eq_landau_n_rt_satz88 x1 x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_satz81e x1 x2 n) (l_e_st_eq_landau_n_rt_rp_5161_t69 zeta m z1 l1 l2 x1 lx1 x2 lx2 i n)))))))))))). Time Defined. (* constant 3011 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t71 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_imp_th1 (l_e_st_eq_landau_n_rt_more x1 x2) (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)) (fun (t:l_e_st_eq_landau_n_rt_more x1 x2) => l_e_st_eq_landau_n_rt_rp_5161_t64 zeta m z1 l1 l2 x1 lx1 x2 lx2 i t) (fun (t:l_not (l_e_st_eq_landau_n_rt_more x1 x2)) => l_e_st_eq_landau_n_rt_rp_5161_t68 zeta m z1 l1 l2 x1 lx1 x2 lx2 i t))))))))))). Time Defined. (* constant 3012 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t72 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_lessis x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_imp_th1 (l_e_st_eq_landau_n_rt_more x1 x2) (l_e_st_eq_landau_n_rt_lessis x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)) (fun (t:l_e_st_eq_landau_n_rt_more x1 x2) => l_e_st_eq_landau_n_rt_rp_5161_t65 zeta m z1 l1 l2 x1 lx1 x2 lx2 i t) (fun (t:l_not (l_e_st_eq_landau_n_rt_more x1 x2)) => l_e_st_eq_landau_n_rt_rp_5161_t70 zeta m z1 l1 l2 x1 lx1 x2 lx2 i t))))))))))). Time Defined. (* constant 3013 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t73 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_lessis x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_imp_th1 (l_e_st_eq_landau_n_rt_more x1 x2) (l_e_st_eq_landau_n_rt_lessis x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)) (fun (t:l_e_st_eq_landau_n_rt_more x1 x2) => l_e_st_eq_landau_n_rt_rp_5161_t66 zeta m z1 l1 l2 x1 lx1 x2 lx2 i t) (fun (t:l_not (l_e_st_eq_landau_n_rt_more x1 x2)) => l_e_st_eq_landau_n_rt_rp_5161_t69 zeta m z1 l1 l2 x1 lx1 x2 lx2 i t))))))))))). Time Defined. (* constant 3014 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t74 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta) (l_e_st_eq_landau_n_rt_rp_5161_t58 zeta) (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t71 zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). Time Defined. (* constant 3015 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t75 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5161_xr1 zeta m z1 l1 l2 x1) (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_st_eq_landau_n_rt_rp_5161_t59 x1 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t72 zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). Time Defined. (* constant 3016 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t76 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5161_xr2 zeta m z1 l1 l2 x1 lx1 x2) (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_st_eq_landau_n_rt_rp_5161_t59 x2 (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t73 zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). Time Defined. (* constant 3017 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t77 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr1 zeta m z1 l1 l2 x1) (l_e_st_eq_landau_n_rt_rp_5161_xr2 zeta m z1 l1 l2 x1 lx1 x2)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts x1 x2)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr1 zeta m z1 l1 l2 x1) (l_e_st_eq_landau_n_rt_rp_5161_xr2 zeta m z1 l1 l2 x1 lx1 x2)) (l_e_st_eq_landau_n_rt_rp_satz154b z1 (l_e_st_eq_landau_n_rt_ts x1 x2) i) (l_e_st_eq_landau_n_rt_rp_satz155c x1 x2))))))))))). Time Defined. (* constant 3018 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t78 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_rp_lessis (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_st_eq_landau_n_rt_rp_islessis1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr1 zeta m z1 l1 l2 x1) (l_e_st_eq_landau_n_rt_rp_5161_xr2 zeta m z1 l1 l2 x1 lx1 x2)) (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)) (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xr1 zeta m z1 l1 l2 x1) (l_e_st_eq_landau_n_rt_rp_5161_xr2 zeta m z1 l1 l2 x1 lx1 x2)) (l_e_st_eq_landau_n_rt_rp_5161_t77 zeta m z1 l1 l2 x1 lx1 x2 lx2 i)) (l_e_st_eq_landau_n_rt_rp_satz149a (l_e_st_eq_landau_n_rt_rp_5161_xr1 zeta m z1 l1 l2 x1) (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_xr2 zeta m z1 l1 l2 x1 lx1 x2) (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t75 zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t76 zeta m z1 l1 l2 x1 lx1 x2 lx2 i)))))))))))). Time Defined. (* constant 3019 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t79 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)) zeta)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) zeta) (l_e_st_eq_landau_n_rt_rp_5161_xm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t74 zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). Time Defined. (* constant 3020 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t80 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) zeta)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_st_eq_landau_n_rt_rp_satz127a (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_xrm zeta m z1 l1 l2 x1 lx1 x2 lx2 i)) zeta (l_e_st_eq_landau_n_rt_rp_5161_t78 zeta m z1 l1 l2 x1 lx1 x2 lx2 i) (l_e_st_eq_landau_n_rt_rp_5161_t79 zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). Time Defined. (* constant 3021 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t81 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_con)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_ec3e23 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) zeta) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) zeta) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) zeta) (l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) zeta) (l_e_st_eq_landau_n_rt_rp_satz122 zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) l1) (l_e_st_eq_landau_n_rt_rp_5161_t80 zeta m z1 l1 l2 x1 lx1 x2 lx2 i))))))))))). Time Defined. (* constant 3022 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t82 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), (forall (x1:l_e_st_eq_landau_n_rt_rat), (forall (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1), (forall (x2:l_e_st_eq_landau_n_rt_rat), (forall (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2), (forall (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)), l_con)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => (fun (x1:l_e_st_eq_landau_n_rt_rat) => (fun (lx1:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x1) => (fun (x2:l_e_st_eq_landau_n_rt_rat) => (fun (lx2:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x2) => (fun (i:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x1 x2)) => l_e_st_eq_landau_n_rt_rp_tsapp (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) z1 (l_e_st_eq_landau_n_rt_rp_5161_t62 zeta m z1 l1 l2) l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) y) => (fun (v:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_5161_t81 zeta m z1 l1 l2 x t y u v))))))))))))))). Time Defined. (* constant 3023 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t82a : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z1:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))), l_con))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z1:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr1 zeta m z1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => l_e_st_eq_landau_n_rt_rp_tsapp (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) z1 (l_e_st_eq_landau_n_rt_rp_5161_t62 zeta m z1 l1 l2) l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) x) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) y) => (fun (v:l_e_st_eq_landau_n_rt_is z1 (l_e_st_eq_landau_n_rt_ts x y)) => l_e_st_eq_landau_n_rt_rp_5161_t82 zeta m z1 l1 l2 x t y u v)))))))))). Time Defined. (* constant 3024 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t83 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), l_con)). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => l_e_st_eq_landau_n_rt_rp_satz159app zeta (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_t61 zeta m) l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_less zeta (l_e_st_eq_landau_n_rt_rp_rpofrt x)) => (fun (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))) => l_e_st_eq_landau_n_rt_rp_5161_t82a zeta m x t u))))). Time Defined. (* constant 3025 *) Definition l_e_st_eq_landau_n_rt_rp_5161_zr2 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt z2))). Time Defined. (* constant 3026 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t84 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) l3))))). Time Defined. (* constant 3027 *) Definition l_e_st_eq_landau_n_rt_rp_5161_yr1 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt y1)))))). Time Defined. (* constant 3028 *) Definition l_e_st_eq_landau_n_rt_rp_5161_yr2 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt y2)))))))). Time Defined. (* constant 3029 *) Definition l_e_st_eq_landau_n_rt_rp_5161_ym : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_rat)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_e_ite (l_e_st_eq_landau_n_rt_less y1 y2) l_e_st_eq_landau_n_rt_rat y1 y2)))))))))). Time Defined. (* constant 3030 *) Definition l_e_st_eq_landau_n_rt_rp_5161_yrm : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_cut)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i))))))))))). Time Defined. (* constant 3031 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t85 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (k:l_e_st_eq_landau_n_rt_less y1 y2), l_e_st_eq_landau_n_rt_is y1 (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (k:l_e_st_eq_landau_n_rt_less y1 y2) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) y1 (l_e_itet (l_e_st_eq_landau_n_rt_less y1 y2) l_e_st_eq_landau_n_rt_rat y1 y2 k)))))))))))). Time Defined. (* constant 3032 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t86 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (k:l_e_st_eq_landau_n_rt_less y1 y2), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (k:l_e_st_eq_landau_n_rt_less y1 y2) => l_e_st_eq_landau_n_rt_rp_satz154b y1 (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_t85 zeta l z2 l3 l4 y1 m1 y2 m2 i k)))))))))))). Time Defined. (* constant 3033 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t87 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (k:l_e_st_eq_landau_n_rt_less y1 y2), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (k:l_e_st_eq_landau_n_rt_less y1 y2) => l_e_st_eq_landau_n_rt_rp_ismore1 (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_t86 zeta l z2 l3 l4 y1 m1 y2 m2 i k) m1))))))))))). Time Defined. (* constant 3034 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t88 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (k:l_e_st_eq_landau_n_rt_less y1 y2), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (k:l_e_st_eq_landau_n_rt_less y1 y2) => l_e_st_eq_landau_n_rt_rp_moreisi2 (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_t86 zeta l z2 l3 l4 y1 m1 y2 m2 i k)))))))))))). Time Defined. (* constant 3035 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t89 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (k:l_e_st_eq_landau_n_rt_less y1 y2), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (k:l_e_st_eq_landau_n_rt_less y1 y2) => l_e_st_eq_landau_n_rt_rp_moreisi1 (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_satz127d (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_satz154c y1 y2 k)) (l_e_st_eq_landau_n_rt_rp_5161_t88 zeta l z2 l3 l4 y1 m1 y2 m2 i k))))))))))))). Time Defined. (* constant 3036 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t90 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)), l_e_st_eq_landau_n_rt_is y2 (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)) => l_e_symis l_e_st_eq_landau_n_rt_rat (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) y2 (l_e_itef (l_e_st_eq_landau_n_rt_less y1 y2) l_e_st_eq_landau_n_rt_rat y1 y2 n)))))))))))). Time Defined. (* constant 3037 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t91 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)) => l_e_st_eq_landau_n_rt_rp_satz154b y2 (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_t90 zeta l z2 l3 l4 y1 m1 y2 m2 i n)))))))))))). Time Defined. (* constant 3038 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t92 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)) => l_e_st_eq_landau_n_rt_rp_ismore1 (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_t91 zeta l z2 l3 l4 y1 m1 y2 m2 i n) m2))))))))))). Time Defined. (* constant 3039 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t93 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)) => l_e_st_eq_landau_n_rt_rp_moreisi2 (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_t91 zeta l z2 l3 l4 y1 m1 y2 m2 i n)))))))))))). Time Defined. (* constant 3040 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t94 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), (forall (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => (fun (n:l_not (l_e_st_eq_landau_n_rt_less y1 y2)) => l_e_st_eq_landau_n_rt_rp_trmoreis (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_t60 y1 y2 (l_e_st_eq_landau_n_rt_satz81f y1 y2 n)) (l_e_st_eq_landau_n_rt_rp_5161_t93 zeta l z2 l3 l4 y1 m1 y2 m2 i n)))))))))))). Time Defined. (* constant 3041 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t95 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_imp_th1 (l_e_st_eq_landau_n_rt_less y1 y2) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (fun (t:l_e_st_eq_landau_n_rt_less y1 y2) => l_e_st_eq_landau_n_rt_rp_5161_t87 zeta l z2 l3 l4 y1 m1 y2 m2 i t) (fun (t:l_not (l_e_st_eq_landau_n_rt_less y1 y2)) => l_e_st_eq_landau_n_rt_rp_5161_t92 zeta l z2 l3 l4 y1 m1 y2 m2 i t))))))))))). Time Defined. (* constant 3042 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t96 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_imp_th1 (l_e_st_eq_landau_n_rt_less y1 y2) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) (fun (t:l_e_st_eq_landau_n_rt_less y1 y2) => l_e_st_eq_landau_n_rt_rp_5161_t88 zeta l z2 l3 l4 y1 m1 y2 m2 i t) (fun (t:l_not (l_e_st_eq_landau_n_rt_less y1 y2)) => l_e_st_eq_landau_n_rt_rp_5161_t94 zeta l z2 l3 l4 y1 m1 y2 m2 i t))))))))))). Time Defined. (* constant 3043 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t97 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_imp_th1 (l_e_st_eq_landau_n_rt_less y1 y2) (l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) (fun (t:l_e_st_eq_landau_n_rt_less y1 y2) => l_e_st_eq_landau_n_rt_rp_5161_t89 zeta l z2 l3 l4 y1 m1 y2 m2 i t) (fun (t:l_not (l_e_st_eq_landau_n_rt_less y1 y2)) => l_e_st_eq_landau_n_rt_rp_5161_t93 zeta l z2 l3 l4 y1 m1 y2 m2 i t))))))))))). Time Defined. (* constant 3044 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t98 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_e_st_eq_landau_n_rt_rp_satz158d (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_moreisi1 (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_t95 zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). Time Defined. (* constant 3045 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t99 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_not (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_imp_th3 (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i)) (l_e_st_eq_landau_n_rt_rp_5161_t98 zeta l z2 l3 l4 y1 m1 y2 m2 i) (fun (t:l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) => l_e_st_eq_landau_n_rt_rp_ine (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta) (l_e_st_eq_landau_n_rt_rp_5161_t58 zeta) (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) t))))))))))). Time Defined. (* constant 3046 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t100 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_not (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) zeta))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) zeta) (l_e_st_eq_landau_n_rt_in (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_sqrtset zeta)) (l_e_st_eq_landau_n_rt_rp_5161_t99 zeta l z2 l3 l4 y1 m1 y2 m2 i) (fun (t:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) zeta) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) zeta) (l_e_st_eq_landau_n_rt_rp_5161_ym zeta l z2 l3 l4 y1 m1 y2 m2 i) t))))))))))). Time Defined. (* constant 3047 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t101 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) zeta)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_e_st_eq_landau_n_rt_rp_satz123f (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) zeta (l_e_st_eq_landau_n_rt_rp_5161_t100 zeta l z2 l3 l4 y1 m1 y2 m2 i))))))))))). Time Defined. (* constant 3048 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t101a : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_e_st_eq_landau_n_rt_rp_satz149 (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_t96 zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_t97 zeta l z2 l3 l4 y1 m1 y2 m2 i))))))))))). Time Defined. (* constant 3049 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t102 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)))))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_e_st_eq_landau_n_rt_rp_ismoreis1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) (l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts y1 y2)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts y1 y2)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2)) (l_e_st_eq_landau_n_rt_rp_satz155c y1 y2)) (l_e_st_eq_landau_n_rt_rp_satz154b (l_e_st_eq_landau_n_rt_ts y1 y2) z2 i)) (l_e_st_eq_landau_n_rt_rp_5161_t101a zeta l z2 l3 l4 y1 m1 y2 m2 i))))))))))). Time Defined. (* constant 3050 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t103 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_e_st_eq_landau_n_rt_rp_trmoreis (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_yrm zeta l z2 l3 l4 y1 m1 y2 m2 i)) zeta (l_e_st_eq_landau_n_rt_rp_5161_t102 zeta l z2 l3 l4 y1 m1 y2 m2 i) (l_e_st_eq_landau_n_rt_rp_5161_t101 zeta l z2 l3 l4 y1 m1 y2 m2 i))))))))))). Time Defined. (* constant 3051 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t104 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), (forall (y1:l_e_st_eq_landau_n_rt_rat), (forall (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (y2:l_e_st_eq_landau_n_rt_rat), (forall (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2), l_con)))))))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => (fun (y1:l_e_st_eq_landau_n_rt_rat) => (fun (m1:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr1 zeta l z2 l3 l4 y1) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y2:l_e_st_eq_landau_n_rt_rat) => (fun (m2:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5161_yr2 zeta l z2 l3 l4 y1 m1 y2) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts y1 y2) z2) => l_e_st_eq_landau_n_rt_rp_satz123c (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta (l_e_st_eq_landau_n_rt_rp_5161_t103 zeta l z2 l3 l4 y1 m1 y2 m2 i) l4)))))))))). Time Defined. (* constant 3052 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t105 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), (forall (z2:l_e_st_eq_landau_n_rt_rat), (forall (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)), (forall (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta), l_con))))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => (fun (z2:l_e_st_eq_landau_n_rt_rat) => (fun (l3:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2)) => (fun (l4:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5161_zr2 zeta l z2) zeta) => l_e_st_eq_landau_n_rt_rp_satz160app (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) z2 (l_e_st_eq_landau_n_rt_rp_5161_t84 zeta l z2 l3 l4) l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt x) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) => (fun (v:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x y) z2) => l_e_st_eq_landau_n_rt_rp_5161_t104 zeta l z2 l3 l4 x t y u v)))))))))). Time Defined. (* constant 3053 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t106 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta), l_con)). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => l_e_st_eq_landau_n_rt_rp_satz159app (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta l l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) => (fun (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) zeta) => l_e_st_eq_landau_n_rt_rp_5161_t105 zeta l x t u))))). Time Defined. (* constant 3054 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t107 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_or3e1 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) (l_e_st_eq_landau_n_rt_rp_satz123a (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) (fun (t:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => l_e_st_eq_landau_n_rt_rp_5161_t83 zeta t) (fun (t:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta)) zeta) => l_e_st_eq_landau_n_rt_rp_5161_t106 zeta t)). Time Defined. (* constant 3055 *) Definition l_e_st_eq_landau_n_rt_rp_5161_t108 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_some (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) zeta)). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_somei l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) zeta) (l_e_st_eq_landau_n_rt_rp_5161_rtc zeta) (l_e_st_eq_landau_n_rt_rp_5161_t107 zeta)). Time Defined. (* constant 3056 *) Definition l_e_st_eq_landau_n_rt_rp_satz161 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_one (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) zeta)). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_onei l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) zeta) (l_e_st_eq_landau_n_rt_rp_5161_t21 zeta) (l_e_st_eq_landau_n_rt_rp_5161_t108 zeta)). Time Defined. (* constant 3057 *) Definition l_e_st_eq_landau_n_rt_rp_irratrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), Prop). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_not (l_e_st_eq_landau_n_rt_rp_ratrp ksi)). Time Defined. (* constant 3058 *) Definition l_e_st_eq_landau_n_5162_t1 : (forall (v:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) v)). exact (fun (v:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl v v) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 v) v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) v) (l_e_st_eq_landau_n_ispl1 v (l_e_st_eq_landau_n_ts l_e_st_eq_landau_n_1 v) v (l_e_st_eq_landau_n_satz28g v)) (l_e_st_eq_landau_n_satz28h l_e_st_eq_landau_n_1 v)). Time Defined. (* constant 3059 *) Definition l_e_st_eq_landau_n_5162_t2 : (forall (v:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_less v (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) v)). exact (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_isless2 (l_e_st_eq_landau_n_pl v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) v) v (l_e_st_eq_landau_n_5162_t1 v) (l_e_st_eq_landau_n_satz18a v v)). Time Defined. (* constant 3060 *) Definition l_e_st_eq_landau_n_5162_t3 : (forall (v:l_e_st_eq_landau_n_nat), (forall (w:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts w w)), l_e_st_eq_landau_n_less v w))). exact (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts w w)) => l_e_st_eq_landau_n_satz10j v w (l_imp_th3 (l_e_st_eq_landau_n_moreis v w) (l_e_st_eq_landau_n_moreis (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts w w)) (l_e_st_eq_landau_n_satz10h (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts w w) l) (fun (t:l_e_st_eq_landau_n_moreis v w) => l_e_st_eq_landau_n_satz36 v w v w t t))))). Time Defined. (* constant 3061 *) Definition l_e_st_eq_landau_n_5162_t4 : (forall (v:l_e_st_eq_landau_n_nat), (forall (w:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) v) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)))). exact (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) v) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts w v)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_disttp1 v w v) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_ts w v) (l_e_st_eq_landau_n_ts v w) (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_comts w v)))). Time Defined. (* constant 3062 *) Definition l_e_st_eq_landau_n_5162_t5 : (forall (v:l_e_st_eq_landau_n_nat), (forall (w:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) (l_e_st_eq_landau_n_pl v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts w w)))). exact (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_nat) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) (l_e_st_eq_landau_n_pl v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v w) (l_e_st_eq_landau_n_ts w w))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts w w)) (l_e_st_eq_landau_n_disttp2 (l_e_st_eq_landau_n_pl v w) v w) (l_e_st_eq_landau_n_ispl12 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) v) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) w) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v w) (l_e_st_eq_landau_n_ts w w)) (l_e_st_eq_landau_n_5162_t4 v w) (l_e_st_eq_landau_n_disttp1 v w w)) (l_e_st_eq_landau_n_asspl2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts v w) (l_e_st_eq_landau_n_ts w w)))). Time Defined. (* constant 3063 *) Definition l_e_st_eq_landau_n_5162_t6 : (forall (v:l_e_st_eq_landau_n_nat), (forall (w:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts v w))))). exact (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v w) (l_e_st_eq_landau_n_ts v w))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts v w))) (l_e_st_eq_landau_n_asspl1 (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v w) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_5162_t1 (l_e_st_eq_landau_n_ts v w))))). Time Defined. (* constant 3064 *) Definition l_e_st_eq_landau_n_5162_nun : (forall (v:l_e_st_eq_landau_n_nat), (forall (w:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) (l_e_st_eq_landau_n_pl v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts v w))) (l_e_st_eq_landau_n_ts w w)))). exact (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) (l_e_st_eq_landau_n_pl v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts w w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts v w))) (l_e_st_eq_landau_n_ts w w)) (l_e_st_eq_landau_n_5162_t5 v w) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_ts v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts v w))) (l_e_st_eq_landau_n_ts w w) (l_e_st_eq_landau_n_5162_t6 v w)))). Time Defined. (* constant 3065 *) Definition l_e_st_eq_landau_n_5162_nun1 : (forall (v:l_e_st_eq_landau_n_nat), (forall (w:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts v w))) (l_e_st_eq_landau_n_ts w w)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) (l_e_st_eq_landau_n_pl v w)))). exact (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl v w) (l_e_st_eq_landau_n_pl v w)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts v v) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts v w))) (l_e_st_eq_landau_n_ts w w)) (l_e_st_eq_landau_n_5162_nun v w))). Time Defined. (* constant 3066 *) Definition l_e_st_eq_landau_n_5162_prop1 : (forall (v:l_e_st_eq_landau_n_nat), (forall (w:l_e_st_eq_landau_n_nat), Prop)). exact (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr w v) (l_e_st_eq_landau_n_fr w v)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1))). Time Defined. (* constant 3067 *) Definition l_e_st_eq_landau_n_5162_prop2 : (forall (v:l_e_st_eq_landau_n_nat), Prop). exact (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_some (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop1 v t)). Time Defined. (* constant 3068 *) Definition l_e_st_eq_landau_n_5162_prop3 : Prop. exact (l_e_st_eq_landau_n_some (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u)). Time Defined. (* constant 3069 *) Definition l_e_st_eq_landau_n_5162_y : (forall (p:l_e_st_eq_landau_n_5162_prop3), l_e_st_eq_landau_n_nat). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => l_e_ind l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u) t) (l_e_st_eq_landau_n_satz27a (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u) p)). Time Defined. (* constant 3070 *) Definition l_e_st_eq_landau_n_5162_t7 : (forall (p:l_e_st_eq_landau_n_5162_prop3), l_e_st_eq_landau_n_min (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u) (l_e_st_eq_landau_n_5162_y p)). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => l_e_oneax l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_min (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u) t) (l_e_st_eq_landau_n_satz27a (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u) p)). Time Defined. (* constant 3071 *) Definition l_e_st_eq_landau_n_5162_t8 : (forall (p:l_e_st_eq_landau_n_5162_prop3), l_e_st_eq_landau_n_lb (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u) (l_e_st_eq_landau_n_5162_y p)). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => l_ande1 (l_e_st_eq_landau_n_lb (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_prop2 (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_t7 p)). Time Defined. (* constant 3072 *) Definition l_e_st_eq_landau_n_5162_t9 : (forall (p:l_e_st_eq_landau_n_5162_prop3), l_e_st_eq_landau_n_5162_prop2 (l_e_st_eq_landau_n_5162_y p)). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => l_ande2 (l_e_st_eq_landau_n_lb (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 u) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_prop2 (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_t7 p)). Time Defined. (* constant 3073 *) Definition l_e_st_eq_landau_n_5162_t10 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => l_e_st_eq_landau_n_treq1 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr x (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_fr x (l_e_st_eq_landau_n_5162_y p))) q (l_e_st_eq_landau_n_tfeq12a x (l_e_st_eq_landau_n_5162_y p) x (l_e_st_eq_landau_n_5162_y p))))). Time Defined. (* constant 3074 *) Definition l_e_st_eq_landau_n_5162_t11 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts x x)))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)))) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts x x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_12isnd (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_5162_t10 p x q) (l_e_st_eq_landau_n_ndis12 (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz28a (l_e_st_eq_landau_n_ts x x))))). Time Defined. (* constant 3075 *) Definition l_e_st_eq_landau_n_5162_t12 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts x x)))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => l_e_st_eq_landau_n_isless2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_t11 p x q) (l_e_st_eq_landau_n_5162_t2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)))))). Time Defined. (* constant 3076 *) Definition l_e_st_eq_landau_n_5162_t13 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => l_e_st_eq_landau_n_isless1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p))) (l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_assts1 (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_t11 p x q)) (l_e_st_eq_landau_n_satz35c (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_refis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)))) (l_e_st_eq_landau_n_5162_t2 (l_e_st_eq_landau_n_5162_y p)))))). Time Defined. (* constant 3077 *) Definition l_e_st_eq_landau_n_5162_t14 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_5162_y p) x))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => l_e_st_eq_landau_n_5162_t3 (l_e_st_eq_landau_n_5162_y p) x (l_e_st_eq_landau_n_5162_t12 p x q)))). Time Defined. (* constant 3078 *) Definition l_e_st_eq_landau_n_5162_t15 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), l_e_st_eq_landau_n_less x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => l_e_st_eq_landau_n_5162_t3 x (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_t13 p x q)))). Time Defined. (* constant 3079 *) Definition l_e_st_eq_landau_n_5162_t16 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => l_e_st_eq_landau_n_isless12 x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) i (l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_t1 (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_5162_t15 p x q)))))). Time Defined. (* constant 3080 *) Definition l_e_st_eq_landau_n_5162_t17 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), l_e_st_eq_landau_n_less u (l_e_st_eq_landau_n_5162_y p)))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => l_e_st_eq_landau_n_satz20f u (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_t16 p x q u i)))))). Time Defined. (* constant 3081 *) Definition l_e_st_eq_landau_n_5162_t18 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl u t) (l_e_st_eq_landau_n_5162_y p)))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t) j))))))). Time Defined. (* constant 3082 *) Definition l_e_st_eq_landau_n_5162_t19 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_ts u u))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ists12 x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u) x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u) i i) (l_e_st_eq_landau_n_5162_nun (l_e_st_eq_landau_n_5162_y p) u)))))))). Time Defined. (* constant 3083 *) Definition l_e_st_eq_landau_n_5162_t20 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t)))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t))) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_5162_t19 p x q u i t j)) (l_e_st_eq_landau_n_asspl1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t))))))))). Time Defined. (* constant 3084 *) Definition l_e_st_eq_landau_n_5162_t21 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts u t))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u) (l_e_st_eq_landau_n_ts u (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts u (l_e_st_eq_landau_n_pl u t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_comts (l_e_st_eq_landau_n_5162_y p) u) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t) u j) (l_e_st_eq_landau_n_disttp2 u u t)))))))). Time Defined. (* constant 3085 *) Definition l_e_st_eq_landau_n_5162_t22 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts u t))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t))) (l_e_st_eq_landau_n_ists2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_5162_t21 p x q u i t j)) (l_e_st_eq_landau_n_disttp2 (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts u t))))))))). Time Defined. (* constant 3086 *) Definition l_e_st_eq_landau_n_5162_t23 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t))) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_5162_t22 p x q u i t j)) (l_e_st_eq_landau_n_asspl2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)))))))))). Time Defined. (* constant 3087 *) Definition l_e_st_eq_landau_n_5162_t24 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t))))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t)))) (l_e_st_eq_landau_n_5162_t20 p x q u i t j) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_5162_t23 p x q u i t j)) (l_e_st_eq_landau_n_asspl1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t)))))))))). Time Defined. (* constant 3088 *) Definition l_e_st_eq_landau_n_5162_t25 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t))) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_pl u t) (l_e_st_eq_landau_n_pl u t)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_asspl2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t))) (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_compl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_5162_nun1 u t) (l_e_st_eq_landau_n_ists12 (l_e_st_eq_landau_n_pl u t) (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t) (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_t18 p x q u i t j) (l_e_st_eq_landau_n_5162_t18 p x q u i t j))))))))). Time Defined. (* constant 3089 *) Definition l_e_st_eq_landau_n_5162_t26 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_5162_t24 p x q u i t j) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_ts t t))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_5162_t25 p x q u i t j)) (l_e_st_eq_landau_n_compl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_asspl2 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)))))))))). Time Defined. (* constant 3090 *) Definition l_e_st_eq_landau_n_5162_t27 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts x x)))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_5162_t1 (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_5162_t11 p x q)))))))). Time Defined. (* constant 3091 *) Definition l_e_st_eq_landau_n_5162_t28 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_5162_t26 p x q u i t j) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p)) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_y p))) (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_5162_t27 p x q u i t j))))))))). Time Defined. (* constant 3092 *) Definition l_e_st_eq_landau_n_5162_t29 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_st_eq_landau_n_satz20e (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts x x) (l_e_st_eq_landau_n_5162_t28 p x q u i t j)))))))). Time Defined. (* constant 3093 *) Definition l_e_st_eq_landau_n_5162_t30 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts u u))))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_tr4is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts u u)))) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_ts t t) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_num (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts u u))) (l_e_st_eq_landau_n_den (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_ndis12 (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts u u)) (l_e_symis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_5162_t29 p x q u i t j)) (l_e_st_eq_landau_n_satz28e (l_e_st_eq_landau_n_ts t t)) (l_e_st_eq_landau_n_12isnd (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts u u) (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1)))))))). Time Defined. (* constant 3094 *) Definition l_e_st_eq_landau_n_5162_t31 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_5162_prop1 u t))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_st_eq_landau_n_treq2 (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr t u) (l_e_st_eq_landau_n_fr t u)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_ts t t) (l_e_st_eq_landau_n_ts u u)) (l_e_st_eq_landau_n_tfeq12a t u t u) (l_e_st_eq_landau_n_5162_t30 p x q u i t j)))))))). Time Defined. (* constant 3095 *) Definition l_e_st_eq_landau_n_5162_t32 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_5162_prop2 u))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_somei l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop1 u v) t (l_e_st_eq_landau_n_5162_t31 p x q u i t j)))))))). Time Defined. (* constant 3096 *) Definition l_e_st_eq_landau_n_5162_t33 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_5162_y p) u))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_st_eq_landau_n_5162_t8 p u (l_e_st_eq_landau_n_5162_t32 p x q u i t j)))))))). Time Defined. (* constant 3097 *) Definition l_e_st_eq_landau_n_5162_t34 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), (forall (t:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)), l_con))))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => (fun (t:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_pl u t)) => l_e_st_eq_landau_n_satz10g (l_e_st_eq_landau_n_5162_y p) u (l_e_st_eq_landau_n_satz12 u (l_e_st_eq_landau_n_5162_y p) (l_e_st_eq_landau_n_5162_t17 p x q u i)) (l_e_st_eq_landau_n_5162_t33 p x q u i t j)))))))). Time Defined. (* constant 3098 *) Definition l_e_st_eq_landau_n_5162_t35 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), (forall (u:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)), l_con))))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_5162_y p) u)) => l_someapp l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_5162_y p) u v) (l_e_st_eq_landau_n_5162_t17 p x q u i) l_con (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_diffprop (l_e_st_eq_landau_n_5162_y p) u v) => l_e_st_eq_landau_n_5162_t34 p x q u i v w))))))). Time Defined. (* constant 3099 *) Definition l_e_st_eq_landau_n_5162_t36 : (forall (p:l_e_st_eq_landau_n_5162_prop3), (forall (x:l_e_st_eq_landau_n_nat), (forall (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x), l_con))). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (q:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) x) => l_someapp l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_diffprop x (l_e_st_eq_landau_n_5162_y p) v) (l_e_st_eq_landau_n_5162_t14 p x q) l_con (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_diffprop x (l_e_st_eq_landau_n_5162_y p) v) => l_e_st_eq_landau_n_5162_t35 p x q v w))))). Time Defined. (* constant 3100 *) Definition l_e_st_eq_landau_n_5162_t37 : (forall (p:l_e_st_eq_landau_n_5162_prop3), l_con). exact (fun (p:l_e_st_eq_landau_n_5162_prop3) => l_someapp l_e_st_eq_landau_n_nat (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) v) (l_e_st_eq_landau_n_5162_t9 p) l_con (fun (v:l_e_st_eq_landau_n_nat) => (fun (w:l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_5162_y p) v) => l_e_st_eq_landau_n_5162_t36 p v w))). Time Defined. (* constant 3101 *) Definition l_e_st_eq_landau_n_rt_5162_t38 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf x x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_ise (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_tf x x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_tict x0 x0 x x xix0 xix0) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1)) i)))). Time Defined. (* constant 3102 *) Definition l_e_st_eq_landau_n_rt_5162_t39 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) x)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_refeq1 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) x (l_e_st_eq_landau_n_fris x))))). Time Defined. (* constant 3103 *) Definition l_e_st_eq_landau_n_rt_5162_t40 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_eq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x))) (l_e_st_eq_landau_n_tf x x))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_eqtf12 (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) x (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) x (l_e_st_eq_landau_n_rt_5162_t39 x0 i x xix0) (l_e_st_eq_landau_n_rt_5162_t39 x0 i x xix0))))). Time Defined. (* constant 3104 *) Definition l_e_st_eq_landau_n_rt_5162_t41 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_den x) (l_e_st_eq_landau_n_num x))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_treq (l_e_st_eq_landau_n_tf (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x)) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_den x))) (l_e_st_eq_landau_n_tf x x) (l_e_st_eq_landau_n_fr (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_5162_t40 x0 i x xix0) (l_e_st_eq_landau_n_rt_5162_t38 x0 i x xix0))))). Time Defined. (* constant 3105 *) Definition l_e_st_eq_landau_n_rt_5162_t42 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_5162_prop2 (l_e_st_eq_landau_n_den x))))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop1 (l_e_st_eq_landau_n_den x) t) (l_e_st_eq_landau_n_num x) (l_e_st_eq_landau_n_rt_5162_t41 x0 i x xix0))))). Time Defined. (* constant 3106 *) Definition l_e_st_eq_landau_n_rt_5162_t43 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_e_st_eq_landau_n_5162_prop3)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_5162_prop2 t) (l_e_st_eq_landau_n_den x) (l_e_st_eq_landau_n_rt_5162_t42 x0 i x xix0))))). Time Defined. (* constant 3107 *) Definition l_e_st_eq_landau_n_rt_5162_t44 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))), (forall (x:l_e_st_eq_landau_n_frac), (forall (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)), l_con)))). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) => (fun (x:l_e_st_eq_landau_n_frac) => (fun (xix0:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_5162_t37 (l_e_st_eq_landau_n_rt_5162_t43 x0 i x xix0))))). Time Defined. (* constant 3108 *) Definition l_e_st_eq_landau_n_rt_5162_t45 : (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))), l_con)). exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts x0 x0) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) => l_e_st_eq_landau_n_rt_ratapp1 x0 l_con (fun (x:l_e_st_eq_landau_n_frac) => (fun (t:l_e_st_eq_landau_n_rt_inf x (l_e_st_eq_landau_n_rt_class x0)) => l_e_st_eq_landau_n_rt_5162_t44 x0 i x t)))). Time Defined. (* constant 3109 *) Definition l_e_st_eq_landau_n_rt_rp_5162_ksi : l_e_st_eq_landau_n_rt_cut. exact (l_e_ind l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_satz161 (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1)))). Time Defined. (* constant 3110 *) Definition l_e_st_eq_landau_n_rt_rp_5162_t46 : l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_5162_ksi l_e_st_eq_landau_n_rt_rp_5162_ksi) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1)). exact (l_e_oneax l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_satz161 (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1)))). Time Defined. (* constant 3111 *) Definition l_e_st_eq_landau_n_rt_rp_5162_x0 : (forall (r:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi), l_e_st_eq_landau_n_rt_rat). exact (fun (r:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi) => l_e_st_eq_landau_n_rt_rp_rtofrp l_e_st_eq_landau_n_rt_rp_5162_ksi r). Time Defined. (* constant 3112 *) Definition l_e_st_eq_landau_n_rt_rp_5162_t47 : (forall (r:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5162_x0 r) (l_e_st_eq_landau_n_rt_rp_5162_x0 r))) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5162_x0 r) (l_e_st_eq_landau_n_rt_rp_5162_x0 r))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5162_x0 r)) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5162_x0 r))) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_5162_ksi l_e_st_eq_landau_n_rt_rp_5162_ksi) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_satz155c (l_e_st_eq_landau_n_rt_rp_5162_x0 r) (l_e_st_eq_landau_n_rt_rp_5162_x0 r)) (l_e_st_eq_landau_n_rt_rp_ists12 (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5162_x0 r)) l_e_st_eq_landau_n_rt_rp_5162_ksi (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_rp_5162_x0 r)) l_e_st_eq_landau_n_rt_rp_5162_ksi (l_e_st_eq_landau_n_rt_rp_isrprt2 l_e_st_eq_landau_n_rt_rp_5162_ksi r) (l_e_st_eq_landau_n_rt_rp_isrprt2 l_e_st_eq_landau_n_rt_rp_5162_ksi r)) l_e_st_eq_landau_n_rt_rp_5162_t46). Time Defined. (* constant 3113 *) Definition l_e_st_eq_landau_n_rt_rp_5162_t48 : (forall (r:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5162_x0 r) (l_e_st_eq_landau_n_rt_rp_5162_x0 r)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi) => l_e_st_eq_landau_n_rt_rp_isrtirp (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_5162_x0 r) (l_e_st_eq_landau_n_rt_rp_5162_x0 r)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_suc l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_5162_t47 r)). Time Defined. (* constant 3114 *) Definition l_e_st_eq_landau_n_rt_rp_5162_t49 : (forall (r:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi), l_con). exact (fun (r:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi) => l_e_st_eq_landau_n_rt_5162_t45 (l_e_st_eq_landau_n_rt_rp_5162_x0 r) (l_e_st_eq_landau_n_rt_rp_5162_t48 r)). Time Defined. (* constant 3115 *) Definition l_e_st_eq_landau_n_rt_rp_satz162 : l_e_st_eq_landau_n_rt_rp_some (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_irratrp a). exact (l_somei l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_irratrp a) l_e_st_eq_landau_n_rt_rp_5162_ksi (fun (t:l_e_st_eq_landau_n_rt_rp_ratrp l_e_st_eq_landau_n_rt_rp_5162_ksi) => l_e_st_eq_landau_n_rt_rp_5162_t49 t)). Time Defined. (* constant 3116 *) Definition l_e_st_eq_landau_n_rt_rp_sqrt : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_ind l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) zeta) (l_e_st_eq_landau_n_rt_rp_satz161 zeta)). Time Defined. (* constant 3117 *) Definition l_e_st_eq_landau_n_rt_rp_thsqrt1 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_sqrt zeta) (l_e_st_eq_landau_n_rt_rp_sqrt zeta)) zeta). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => l_e_oneax l_e_st_eq_landau_n_rt_cut (fun (a:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts a a) zeta) (l_e_st_eq_landau_n_rt_rp_satz161 zeta)). Time Defined. (* constant 3118 *) Definition l_e_st_eq_landau_n_rt_rp_thsqrt2 : (forall (zeta:l_e_st_eq_landau_n_rt_cut), (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi ksi) zeta), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_sqrt zeta)))). exact (fun (zeta:l_e_st_eq_landau_n_rt_cut) => (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi ksi) zeta) => l_e_st_eq_landau_n_rt_rp_5161_t20 zeta ksi (l_e_st_eq_landau_n_rt_rp_sqrt zeta) i (l_e_st_eq_landau_n_rt_rp_thsqrt1 zeta)))). Time Defined. (* constant 3119 *) Definition l_e_st_eq_landau_n_rt_rp_issqrt : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_sqrt ksi) (l_e_st_eq_landau_n_rt_rp_sqrt eta)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is ksi eta) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_sqrt t) ksi eta i))). Time Defined. (* constant 3120 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_x : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_nat)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rp_ntofrp ksi nx)))). Time Defined. (* constant 3121 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_y : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_nat)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rp_ntofrp eta ny)))). Time Defined. (* constant 3122 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t1 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is ksi (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rp_isrpnt1 ksi nx)))). Time Defined. (* constant 3123 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t2 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is eta (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rp_isrpnt1 eta ny)))). Time Defined. (* constant 3124 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t3 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi eta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rp_ispl12 ksi (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) eta (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t1 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t2 ksi nx eta ny))))). Time Defined. (* constant 3125 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_x0 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rat)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny))))). Time Defined. (* constant 3126 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_y0 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rat)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))))). Time Defined. (* constant 3127 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t4 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_natrti (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny))))). Time Defined. (* constant 3128 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t5 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_natrti (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))))). Time Defined. (* constant 3129 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t6 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_satz155a (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)))))). Time Defined. (* constant 3130 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t7 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_satz112d (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t4 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t5 ksi nx eta ny))))). Time Defined. (* constant 3131 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_xpy : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_nat)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t7 ksi nx eta ny))))). Time Defined. (* constant 3132 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t8 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_iiia_xpy ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_isrtn1 (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t7 ksi nx eta ny))))). Time Defined. (* constant 3133 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t9 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xpy ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rp_isrterp (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_iiia_xpy ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t8 ksi nx eta ny))))). Time Defined. (* constant 3134 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t10 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi eta) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xpy ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl ksi eta) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_pl (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xpy ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t3 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t6 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t9 ksi nx eta ny))))). Time Defined. (* constant 3135 *) Definition l_e_st_eq_landau_n_rt_rp_natpl : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_pl ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl ksi eta) (l_e_st_eq_landau_n_rt_rp_rpofnt t)) (l_e_st_eq_landau_n_rt_rp_iiia_xpy ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t10 ksi nx eta ny))))). Time Defined. (* constant 3136 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t11 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rp_ists12 ksi (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) eta (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t1 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t2 ksi nx eta ny))))). Time Defined. (* constant 3137 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t12 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_satz155c (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)))))). Time Defined. (* constant 3138 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t13 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_satz112f (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t4 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t5 ksi nx eta ny))))). Time Defined. (* constant 3139 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_xty : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_nat)))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t13 ksi nx eta ny))))). Time Defined. (* constant 3140 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t14 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_iiia_xty ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_isrtn1 (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t13 ksi nx eta ny))))). Time Defined. (* constant 3141 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t15 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xty ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_st_eq_landau_n_rt_rp_isrterp (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_iiia_xty ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t14 ksi nx eta ny))))). Time Defined. (* constant 3142 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t16 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xty ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_ts (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny))) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xty ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t11 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t12 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t15 ksi nx eta ny))))). Time Defined. (* constant 3143 *) Definition l_e_st_eq_landau_n_rt_rp_natts : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_ts ksi eta))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts ksi eta) (l_e_st_eq_landau_n_rt_rp_rpofnt t)) (l_e_st_eq_landau_n_rt_rp_iiia_xty ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t16 ksi nx eta ny))))). Time Defined. (* constant 3144 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t17 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_ismore12 ksi (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) eta (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_iiia_t1 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t2 ksi nx eta ny) m))))). Time Defined. (* constant 3145 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t18 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_satz154d (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t17 ksi nx eta ny m)))))). Time Defined. (* constant 3146 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t20 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_satz154a (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_ismn12 ksi (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) eta (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)) m (l_e_st_eq_landau_n_rt_rp_satz154a (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m)) (l_e_st_eq_landau_n_rt_rp_iiia_t1 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t2 ksi nx eta ny)))))). Time Defined. (* constant 3147 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t21 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_satz154a (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m))) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m)))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m))) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_satz154a (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m))) (l_e_st_eq_landau_n_rt_rp_satz155b (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m))))))). Time Defined. (* constant 3148 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t22 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_natrt (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_satz112g (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t4 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t5 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m)))))). Time Defined. (* constant 3149 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_xmy : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_nat))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_nofrt (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m)) (l_e_st_eq_landau_n_rt_rp_iiia_t22 ksi nx eta ny m)))))). Time Defined. (* constant 3150 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t23 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_is (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_iiia_xmy ksi nx eta ny m))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_isrtn1 (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m)) (l_e_st_eq_landau_n_rt_rp_iiia_t22 ksi nx eta ny m)))))). Time Defined. (* constant 3151 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t24 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m))) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xmy ksi nx eta ny m))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_st_eq_landau_n_rt_rp_isrterp (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m)) (l_e_st_eq_landau_n_rt_rtofn (l_e_st_eq_landau_n_rt_rp_iiia_xmy ksi nx eta ny m)) (l_e_st_eq_landau_n_rt_rp_iiia_t23 ksi nx eta ny m)))))). Time Defined. (* constant 3152 *) Definition l_e_st_eq_landau_n_rt_rp_iiia_t25 : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xmy ksi nx eta ny m))))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_x ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_y ksi nx eta ny)) (l_e_st_eq_landau_n_rt_rp_satz154a (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m))) (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn (l_e_st_eq_landau_n_rt_rp_iiia_x0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_y0 ksi nx eta ny) (l_e_st_eq_landau_n_rt_rp_iiia_t18 ksi nx eta ny m))) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_iiia_xmy ksi nx eta ny m)) (l_e_st_eq_landau_n_rt_rp_iiia_t20 ksi nx eta ny m) (l_e_st_eq_landau_n_rt_rp_iiia_t21 ksi nx eta ny m) (l_e_st_eq_landau_n_rt_rp_iiia_t24 ksi nx eta ny m)))))). Time Defined. (* constant 3153 *) Definition l_e_st_eq_landau_n_rt_rp_natmn : (forall (ksi:l_e_st_eq_landau_n_rt_cut), (forall (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi), (forall (eta:l_e_st_eq_landau_n_rt_cut), (forall (ny:l_e_st_eq_landau_n_rt_rp_natrp eta), (forall (m:l_e_st_eq_landau_n_rt_rp_more ksi eta), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_mn ksi eta m)))))). exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => (fun (nx:l_e_st_eq_landau_n_rt_rp_natrp ksi) => (fun (eta:l_e_st_eq_landau_n_rt_cut) => (fun (ny:l_e_st_eq_landau_n_rt_rp_natrp eta) => (fun (m:l_e_st_eq_landau_n_rt_rp_more ksi eta) => l_somei l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_mn ksi eta m) (l_e_st_eq_landau_n_rt_rp_rpofnt t)) (l_e_st_eq_landau_n_rt_rp_iiia_xmy ksi nx eta ny m) (l_e_st_eq_landau_n_rt_rp_iiia_t25 ksi nx eta ny m)))))). Time Defined. (* constant 3154 *) Definition l_e_st_eq_landau_n_rt_rp_3pl13 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl q r)) (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_pl q p))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl q r)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl q r) p) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl r q) p) (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_pl q p)) (l_e_st_eq_landau_n_rt_rp_compl p (l_e_st_eq_landau_n_rt_rp_pl q r)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl q r) (l_e_st_eq_landau_n_rt_rp_pl r q) p (l_e_st_eq_landau_n_rt_rp_compl q r)) (l_e_st_eq_landau_n_rt_rp_asspl1 r q p)))). Time Defined. (* constant 3155 *) Definition l_e_st_eq_landau_n_rt_rp_4pl24 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p q) (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p s) (l_e_st_eq_landau_n_rt_rp_pl r q)))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p q) (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl q (l_e_st_eq_landau_n_rt_rp_pl r s))) (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl s (l_e_st_eq_landau_n_rt_rp_pl r q))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p s) (l_e_st_eq_landau_n_rt_rp_pl r q)) (l_e_st_eq_landau_n_rt_rp_asspl1 p q (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl q (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_pl s (l_e_st_eq_landau_n_rt_rp_pl r q)) p (l_e_st_eq_landau_n_rt_rp_3pl13 q r s)) (l_e_st_eq_landau_n_rt_rp_asspl2 p s (l_e_st_eq_landau_n_rt_rp_pl r q)))))). Time Defined. (* constant 3156 *) Definition l_e_st_eq_landau_n_rt_rp_3pl12 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl q r)) (l_e_st_eq_landau_n_rt_rp_pl q (l_e_st_eq_landau_n_rt_rp_pl p r))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl q r)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p q) r) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl q p) r) (l_e_st_eq_landau_n_rt_rp_pl q (l_e_st_eq_landau_n_rt_rp_pl p r)) (l_e_st_eq_landau_n_rt_rp_asspl2 p q r) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl p q) (l_e_st_eq_landau_n_rt_rp_pl q p) r (l_e_st_eq_landau_n_rt_rp_compl p q)) (l_e_st_eq_landau_n_rt_rp_asspl1 q p r)))). Time Defined. (* constant 3157 *) Definition l_e_st_eq_landau_n_rt_rp_4pl23 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p q) (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p r) (l_e_st_eq_landau_n_rt_rp_pl q s)))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p q) (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl q (l_e_st_eq_landau_n_rt_rp_pl r s))) (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_pl q s))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p r) (l_e_st_eq_landau_n_rt_rp_pl q s)) (l_e_st_eq_landau_n_rt_rp_asspl1 p q (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl q (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_pl q s)) p (l_e_st_eq_landau_n_rt_rp_3pl12 q r s)) (l_e_st_eq_landau_n_rt_rp_asspl2 p r (l_e_st_eq_landau_n_rt_rp_pl q s)))))). Time Defined. (* constant 3158 *) Definition l_e_st_eq_landau_n_rt_rp_3pl23 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p q) r) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p r) q)))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p q) r) (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl q r)) (l_e_st_eq_landau_n_rt_rp_pl p (l_e_st_eq_landau_n_rt_rp_pl r q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl p r) q) (l_e_st_eq_landau_n_rt_rp_asspl1 p q r) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl q r) (l_e_st_eq_landau_n_rt_rp_pl r q) p (l_e_st_eq_landau_n_rt_rp_compl q r)) (l_e_st_eq_landau_n_rt_rp_asspl2 p r q)))). Time Defined. (* constant 3159 *) Definition l_e_st_eq_landau_n_rt_rp_a2isapa : (forall (p:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl p p)). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts p l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl p p) (l_e_st_eq_landau_n_rt_rp_disttp2 p l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts p l_e_st_eq_landau_n_rt_rp_1rp) p (l_e_st_eq_landau_n_rt_rp_ts p l_e_st_eq_landau_n_rt_rp_1rp) p (l_e_st_eq_landau_n_rt_rp_satz151 p) (l_e_st_eq_landau_n_rt_rp_satz151 p))). Time Defined. (* constant 3160 *) Definition l_e_st_eq_landau_n_rt_rp_dif : Type. exact (l_e_st_eq_landau_n_pair1type l_e_st_eq_landau_n_rt_cut). Time Defined. (* constant 3161 *) Definition l_e_st_eq_landau_n_rt_rp_df : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_dif)). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_pair1 l_e_st_eq_landau_n_rt_cut a1 a2)). Time Defined. (* constant 3162 *) Definition l_e_st_eq_landau_n_rt_rp_stm : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_first1 l_e_st_eq_landau_n_rt_cut a). Time Defined. (* constant 3163 *) Definition l_e_st_eq_landau_n_rt_rp_std : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_second1 l_e_st_eq_landau_n_rt_cut a). Time Defined. (* constant 3164 *) Definition l_e_st_eq_landau_n_rt_rp_stmis : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a1)). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_first1is1 l_e_st_eq_landau_n_rt_cut a1 a2)). Time Defined. (* constant 3165 *) Definition l_e_st_eq_landau_n_rt_rp_isstm : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is a1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_first1is2 l_e_st_eq_landau_n_rt_cut a1 a2)). Time Defined. (* constant 3166 *) Definition l_e_st_eq_landau_n_rt_rp_stdis : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2)). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_second1is1 l_e_st_eq_landau_n_rt_cut a1 a2)). Time Defined. (* constant 3167 *) Definition l_e_st_eq_landau_n_rt_rp_isstd : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is a2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_second1is2 l_e_st_eq_landau_n_rt_cut a1 a2)). Time Defined. (* constant 3168 *) Definition l_e_st_eq_landau_n_rt_rp_1a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_stm a). Time Defined. (* constant 3169 *) Definition l_e_st_eq_landau_n_rt_rp_2a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_std a). Time Defined. (* constant 3170 *) Definition l_e_st_eq_landau_n_rt_rp_dfis : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) a). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_pair1is1 l_e_st_eq_landau_n_rt_cut a). Time Defined. (* constant 3171 *) Definition l_e_st_eq_landau_n_rt_rp_isdf : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_is l_e_st_eq_landau_n_rt_rp_dif a (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_pair1is2 l_e_st_eq_landau_n_rt_cut a). Time Defined. (* constant 3172 *) Definition l_e_st_eq_landau_n_rt_rp_12issmsd : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ispl12 a1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) b2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_isstm a1 a2) (l_e_st_eq_landau_n_rt_rp_isstd b1 b2))))). Time Defined. (* constant 3173 *) Definition l_e_st_eq_landau_n_rt_rp_smsdis12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_pl a1 b2))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_12issmsd a1 a2 b1 b2))))). Time Defined. (* constant 3174 *) Definition l_e_st_eq_landau_n_rt_rp_1sdissmsd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ispl1 r1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_isstm r1 r2)))). Time Defined. (* constant 3175 *) Definition l_e_st_eq_landau_n_rt_rp_smsdis1sd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_1sdissmsd a r1 r2)))). Time Defined. (* constant 3176 *) Definition l_e_st_eq_landau_n_rt_rp_sm2issmsd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ispl2 r2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_isstd r1 r2)))). Time Defined. (* constant 3177 *) Definition l_e_st_eq_landau_n_rt_rp_smsdissm2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2))) (l_e_st_eq_landau_n_rt_rp_sm2issmsd a r1 r2)))). Time Defined. (* constant 3178 *) Definition l_e_st_eq_landau_n_rt_rp_issm : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is a1 r), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df r a2))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is a1 r) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_rp_dif (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_df t a2) a1 r i)))). Time Defined. (* constant 3179 *) Definition l_e_st_eq_landau_n_rt_rp_issd : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is a2 r), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df a1 r))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is a2 r) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_rp_dif (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_df a1 t) a2 r i)))). Time Defined. (* constant 3180 *) Definition l_e_st_eq_landau_n_rt_rp_issmsd : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is a1 b1), (forall (j:l_e_st_eq_landau_n_rt_rp_is a2 b2), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2))))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is a1 b1) => (fun (j:l_e_st_eq_landau_n_rt_rp_is a2 b2) => l_e_tris l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2) (l_e_st_eq_landau_n_rt_rp_issm a1 a2 b1 i) (l_e_st_eq_landau_n_rt_rp_issd b1 a2 b2 j))))))). Time Defined. (* constant 3181 *) Definition l_e_st_eq_landau_n_rt_rp_1b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_stm b)). Time Defined. (* constant 3182 *) Definition l_e_st_eq_landau_n_rt_rp_2b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_std b)). Time Defined. (* constant 3183 *) Definition l_e_st_eq_landau_n_rt_rp_eq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3184 *) Definition l_e_st_eq_landau_n_rt_rp_eqi12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_smsdis12 a1 a2 b1 b2) i (l_e_st_eq_landau_n_rt_rp_12issmsd b1 b2 a1 a2)))))). Time Defined. (* constant 3185 *) Definition l_e_st_eq_landau_n_rt_rp_eqi1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a))), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_df r1 r2))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a))) => l_e_isp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq x (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) a (l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) r1 r2 i) (l_e_st_eq_landau_n_rt_rp_dfis a))))). Time Defined. (* constant 3186 *) Definition l_e_st_eq_landau_n_rt_rp_eqi2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df r1 r2) a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2)) => l_e_isp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df r1 r2) x) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) a (l_e_st_eq_landau_n_rt_rp_eqi12 r1 r2 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) i) (l_e_st_eq_landau_n_rt_rp_dfis a))))). Time Defined. (* constant 3187 *) Definition l_e_st_eq_landau_n_rt_rp_eqe12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_pl b1 a2) (l_e_st_eq_landau_n_rt_rp_12issmsd a1 a2 b1 b2) e (l_e_st_eq_landau_n_rt_rp_smsdis12 b1 b2 a1 a2)))))). Time Defined. (* constant 3188 *) Definition l_e_st_eq_landau_n_rt_rp_satzd163 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq a a). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a))). Time Defined. (* constant 3189 *) Definition l_e_st_eq_landau_n_rt_rp_refeq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq a a). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satzd163 a). Time Defined. (* constant 3190 *) Definition l_e_st_eq_landau_n_rt_rp_refeq1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_dif a b), l_e_st_eq_landau_n_rt_rp_eq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_dif a b) => l_e_isp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq a x) a b (l_e_st_eq_landau_n_rt_rp_refeq a) i))). Time Defined. (* constant 3191 *) Definition l_e_st_eq_landau_n_rt_rp_refeq2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_dif a b), l_e_st_eq_landau_n_rt_rp_eq b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_dif a b) => l_e_isp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq x a) a b (l_e_st_eq_landau_n_rt_rp_refeq a) i))). Time Defined. (* constant 3192 *) Definition l_e_st_eq_landau_n_rt_rp_eqsmsd : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is a1 b1), (forall (j:l_e_st_eq_landau_n_rt_rp_is a2 b2), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2))))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is a1 b1) => (fun (j:l_e_st_eq_landau_n_rt_rp_is a2 b2) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2) (l_e_st_eq_landau_n_rt_rp_issmsd a1 a2 b1 b2 i j))))))). Time Defined. (* constant 3193 *) Definition l_e_st_eq_landau_n_rt_rp_eqsm : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is a1 r), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df r a2))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is a1 r) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df r a2) (l_e_st_eq_landau_n_rt_rp_issm a1 a2 r i))))). Time Defined. (* constant 3194 *) Definition l_e_st_eq_landau_n_rt_rp_eqsd : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is a2 r), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df a1 r))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is a2 r) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df a1 r) (l_e_st_eq_landau_n_rt_rp_issd a1 a2 r i))))). Time Defined. (* constant 3195 *) Definition l_e_st_eq_landau_n_rt_rp_satzd164 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) e))). Time Defined. (* constant 3196 *) Definition l_e_st_eq_landau_n_rt_rp_symeq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_satzd164 a b e))). Time Defined. (* constant 3197 *) Definition l_e_st_eq_landau_n_rt_rp_1c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_stm c))). Time Defined. (* constant 3198 *) Definition l_e_st_eq_landau_n_rt_rp_2c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_std c))). Time Defined. (* constant 3199 *) Definition l_e_st_eq_landau_n_rt_rp_1d165_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq b c), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq b c) => l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b)) e f))))). Time Defined. (* constant 3200 *) Definition l_e_st_eq_landau_n_rt_rp_1d165_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq b c), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq b c) => l_e_tr4is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_1d165_t1 a b c e f) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))))))). Time Defined. (* constant 3201 *) Definition l_e_st_eq_landau_n_rt_rp_satzd165 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq b c), l_e_st_eq_landau_n_rt_rp_eq a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq b c) => l_e_st_eq_landau_n_rt_rp_satz136b (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_1d165_t2 a b c e f)))))). Time Defined. (* constant 3202 *) Definition l_e_st_eq_landau_n_rt_rp_treq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq b c), l_e_st_eq_landau_n_rt_rp_eq a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq b c) => l_e_st_eq_landau_n_rt_rp_satzd165 a b c e f))))). Time Defined. (* constant 3203 *) Definition l_e_st_eq_landau_n_rt_rp_treq1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq c a), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c b), l_e_st_eq_landau_n_rt_rp_eq a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq c a) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c b) => l_e_st_eq_landau_n_rt_rp_treq a c b (l_e_st_eq_landau_n_rt_rp_symeq c a e) f))))). Time Defined. (* constant 3204 *) Definition l_e_st_eq_landau_n_rt_rp_treq2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq b c), l_e_st_eq_landau_n_rt_rp_eq a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq b c) => l_e_st_eq_landau_n_rt_rp_treq a c b e (l_e_st_eq_landau_n_rt_rp_symeq b c f)))))). Time Defined. (* constant 3205 *) Definition l_e_st_eq_landau_n_rt_rp_tr3eq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e1:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (e2:l_e_st_eq_landau_n_rt_rp_eq b c), (forall (e3:l_e_st_eq_landau_n_rt_rp_eq c d), l_e_st_eq_landau_n_rt_rp_eq a d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e1:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (e2:l_e_st_eq_landau_n_rt_rp_eq b c) => (fun (e3:l_e_st_eq_landau_n_rt_rp_eq c d) => l_e_st_eq_landau_n_rt_rp_treq a b d e1 (l_e_st_eq_landau_n_rt_rp_treq b c d e2 e3)))))))). Time Defined. (* constant 3206 *) Definition l_e_st_eq_landau_n_rt_rp_tr4eq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_dif), (forall (e1:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (e2:l_e_st_eq_landau_n_rt_rp_eq b c), (forall (e3:l_e_st_eq_landau_n_rt_rp_eq c d), (forall (e4:l_e_st_eq_landau_n_rt_rp_eq d e), l_e_st_eq_landau_n_rt_rp_eq a e))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e1:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (e2:l_e_st_eq_landau_n_rt_rp_eq b c) => (fun (e3:l_e_st_eq_landau_n_rt_rp_eq c d) => (fun (e4:l_e_st_eq_landau_n_rt_rp_eq d e) => l_e_st_eq_landau_n_rt_rp_tr3eq a b c e e1 e2 (l_e_st_eq_landau_n_rt_rp_treq c d e e3 e4)))))))))). Time Defined. (* constant 3207 *) Definition l_e_st_eq_landau_n_rt_rp_posd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)). Time Defined. (* constant 3208 *) Definition l_e_st_eq_landau_n_rt_rp_zero : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)). Time Defined. (* constant 3209 *) Definition l_e_st_eq_landau_n_rt_rp_negd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)). Time Defined. (* constant 3210 *) Definition l_e_st_eq_landau_n_rt_rp_posdi : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more a1 a2), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_df a1 a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more a1 a2) => l_e_st_eq_landau_n_rt_rp_ismore12 a1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_isstm a1 a2) (l_e_st_eq_landau_n_rt_rp_isstd a1 a2) m))). Time Defined. (* constant 3211 *) Definition l_e_st_eq_landau_n_rt_rp_zeroi : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is a1 a2), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_df a1 a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is a1 a2) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a1 a2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stmis a1 a2) i (l_e_st_eq_landau_n_rt_rp_isstd a1 a2)))). Time Defined. (* constant 3212 *) Definition l_e_st_eq_landau_n_rt_rp_negdi : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less a1 a2), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_df a1 a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less a1 a2) => l_e_st_eq_landau_n_rt_rp_isless12 a1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_isstm a1 a2) (l_e_st_eq_landau_n_rt_rp_isstd a1 a2) l))). Time Defined. (* constant 3213 *) Definition l_e_st_eq_landau_n_rt_rp_axrde : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_ec3 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)). Time Defined. (* constant 3214 *) Definition l_e_st_eq_landau_n_rt_rp_axrdo : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_or3 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satz123a (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)). Time Defined. (* constant 3215 *) Definition l_e_st_eq_landau_n_rt_rp_axrd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_orec3 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_orec3i (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_axrdo a) (l_e_st_eq_landau_n_rt_rp_axrde a)). Time Defined. (* constant 3216 *) Definition l_e_st_eq_landau_n_rt_rp_rappd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:Prop), (forall (p1:(forall (t:l_e_st_eq_landau_n_rt_rp_posd a), p)), (forall (p2:(forall (t:l_e_st_eq_landau_n_rt_rp_zero a), p)), (forall (p3:(forall (t:l_e_st_eq_landau_n_rt_rp_negd a), p)), p))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:Prop) => (fun (p1:(forall (t:l_e_st_eq_landau_n_rt_rp_posd a), p)) => (fun (p2:(forall (t:l_e_st_eq_landau_n_rt_rp_zero a), p)) => (fun (p3:(forall (t:l_e_st_eq_landau_n_rt_rp_negd a), p)) => l_or3app (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) p (l_e_st_eq_landau_n_rt_rp_axrdo a) p2 p1 p3))))). Time Defined. (* constant 3217 *) Definition l_e_st_eq_landau_n_rt_rp_pnot0d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_not (l_e_st_eq_landau_n_rt_rp_zero a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_ec3e21 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_axrde a) p)). Time Defined. (* constant 3218 *) Definition l_e_st_eq_landau_n_rt_rp_pnotnd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_not (l_e_st_eq_landau_n_rt_rp_negd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_ec3e23 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_axrde a) p)). Time Defined. (* constant 3219 *) Definition l_e_st_eq_landau_n_rt_rp_0notpd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_not (l_e_st_eq_landau_n_rt_rp_posd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_ec3e12 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_axrde a) z)). Time Defined. (* constant 3220 *) Definition l_e_st_eq_landau_n_rt_rp_0notnd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_not (l_e_st_eq_landau_n_rt_rp_negd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_ec3e13 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_axrde a) z)). Time Defined. (* constant 3221 *) Definition l_e_st_eq_landau_n_rt_rp_nnotpd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_not (l_e_st_eq_landau_n_rt_rp_posd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_ec3e32 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_axrde a) n)). Time Defined. (* constant 3222 *) Definition l_e_st_eq_landau_n_rt_rp_nnot0d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_not (l_e_st_eq_landau_n_rt_rp_zero a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_ec3e31 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_axrde a) n)). Time Defined. (* constant 3223 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) e (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_satz135a (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) p))))). Time Defined. (* constant 3224 *) Definition l_e_st_eq_landau_n_rt_rp_eqposd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_posd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_satz136a (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv1d_t1 a b e p))))). Time Defined. (* constant 3225 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_symeq a b e) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) z) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))))). Time Defined. (* constant 3226 *) Definition l_e_st_eq_landau_n_rt_rp_eqzero : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_zero b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_satz136b (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv1d_t2 a b e z))))). Time Defined. (* constant 3227 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) e (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_satz135c (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) n))))). Time Defined. (* constant 3228 *) Definition l_e_st_eq_landau_n_rt_rp_eqnegd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_negd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_satz136c (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv1d_t3 a b e n))))). Time Defined. (* constant 3229 *) Definition l_e_st_eq_landau_n_rt_rp_zeroeq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), (forall (y:l_e_st_eq_landau_n_rt_rp_zero b), l_e_st_eq_landau_n_rt_rp_eq a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => (fun (y:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b) z (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) y)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)))))). Time Defined. (* constant 3230 *) Definition l_e_st_eq_landau_n_rt_rp_pdofrp : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_dif). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp). Time Defined. (* constant 3231 *) Definition l_e_st_eq_landau_n_rt_rp_ndofrp : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_dif). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp)). Time Defined. (* constant 3232 *) Definition l_e_st_eq_landau_n_rt_rp_isrpepd : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is r s), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is r s) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s) (l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pdofrp x) r s i)))). Time Defined. (* constant 3233 *) Definition l_e_st_eq_landau_n_rt_rp_isrpend : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is r s), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_ndofrp s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is r s) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_ndofrp s) (l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ndofrp x) r s i)))). Time Defined. (* constant 3234 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t4 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) => l_e_st_eq_landau_n_rt_rp_satz136b (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_eqe12 (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp e)))). Time Defined. (* constant 3235 *) Definition l_e_st_eq_landau_n_rt_rp_isrpipd : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)), l_e_st_eq_landau_n_rt_rp_is r s))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) => l_e_st_eq_landau_n_rt_rp_satz136b r s l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_iv1d_t4 r s e)))). Time Defined. (* constant 3236 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t5 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_ndofrp s)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_ndofrp s)) => l_e_st_eq_landau_n_rt_rp_satz136e (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_eqe12 l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) e)))). Time Defined. (* constant 3237 *) Definition l_e_st_eq_landau_n_rt_rp_isrpind : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_ndofrp s)), l_e_st_eq_landau_n_rt_rp_is r s))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_ndofrp s)) => l_e_symis l_e_st_eq_landau_n_rt_cut s r (l_e_st_eq_landau_n_rt_rp_satz136b s r l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_iv1d_t5 r s e))))). Time Defined. (* constant 3238 *) Definition l_e_st_eq_landau_n_rt_rp_posdirp : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_posdi (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_ismore1 (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp r) (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_compl l_e_st_eq_landau_n_rt_rp_1rp r) (l_e_st_eq_landau_n_rt_rp_satz133 l_e_st_eq_landau_n_rt_rp_1rp r))). Time Defined. (* constant 3239 *) Definition l_e_st_eq_landau_n_rt_rp_negdirp : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_ndofrp r)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_negdi l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp r) (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_compl l_e_st_eq_landau_n_rt_rp_1rp r) (l_e_st_eq_landau_n_rt_rp_satz133a l_e_st_eq_landau_n_rt_rp_1rp r))). Time Defined. (* constant 3240 *) Definition l_e_st_eq_landau_n_rt_rp_rpofpd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_cut)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) p)). Time Defined. (* constant 3241 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofpd a p) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_2a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_tr4is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofpd a p) (l_e_st_eq_landau_n_rt_rp_2a a)) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofpd a p) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofpd a p) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofpd a p) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofpd a p) (l_e_st_eq_landau_n_rt_rp_2a a)) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_satz140f (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) p)) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_rpofpd a p) (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_rpofpd a p) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_asspl2 (l_e_st_eq_landau_n_rt_rp_rpofpd a p) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3242 *) Definition l_e_st_eq_landau_n_rt_rp_eqpdrp1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd a p)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_eqi1 a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofpd a p) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_iv1d_t6 a p))). Time Defined. (* constant 3243 *) Definition l_e_st_eq_landau_n_rt_rp_eqpdrp2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd a p)) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_symeq a (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd a p)) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 a p))). Time Defined. (* constant 3244 *) Definition l_e_st_eq_landau_n_rt_rp_rpofnd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_cut)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) n))). Time Defined. (* constant 3245 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnd a n) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_2a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnd a n) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_rpofnd a n)) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_asspl2 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_rpofnd a n) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_rpofnd a n)) (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_satz140c (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) n))) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp))). Time Defined. (* constant 3246 *) Definition l_e_st_eq_landau_n_rt_rp_eqndrp1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd a n)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_eqi1 a l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnd a n) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_iv1d_t7 a n))). Time Defined. (* constant 3247 *) Definition l_e_st_eq_landau_n_rt_rp_eqndrp2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd a n)) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_symeq a (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd a n)) (l_e_st_eq_landau_n_rt_rp_eqndrp1 a n))). Time Defined. (* constant 3248 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t8 : (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd h), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (q:l_e_st_eq_landau_n_rt_rp_posd k), (forall (e:l_e_st_eq_landau_n_rt_rp_eq h k), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd h p)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd k q))))))). exact (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd h) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd k) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq h k) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd h p)) h k (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd k q)) (l_e_st_eq_landau_n_rt_rp_eqpdrp2 h p) e (l_e_st_eq_landau_n_rt_rp_eqpdrp1 k q)))))). Time Defined. (* constant 3249 *) Definition l_e_st_eq_landau_n_rt_rp_eqpderp : (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd h), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (q:l_e_st_eq_landau_n_rt_rp_posd k), (forall (e:l_e_st_eq_landau_n_rt_rp_eq h k), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd h p) (l_e_st_eq_landau_n_rt_rp_rpofpd k q)))))). exact (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd h) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd k) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq h k) => l_e_st_eq_landau_n_rt_rp_isrpipd (l_e_st_eq_landau_n_rt_rp_rpofpd h p) (l_e_st_eq_landau_n_rt_rp_rpofpd k q) (l_e_st_eq_landau_n_rt_rp_iv1d_t8 h p k q e)))))). Time Defined. (* constant 3250 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t9 : (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd h), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (q:l_e_st_eq_landau_n_rt_rp_posd k), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd h p) (l_e_st_eq_landau_n_rt_rp_rpofpd k q)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd h p)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd k q))))))). exact (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd h) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd k) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd h p) (l_e_st_eq_landau_n_rt_rp_rpofpd k q)) => l_e_st_eq_landau_n_rt_rp_isrpepd (l_e_st_eq_landau_n_rt_rp_rpofpd h p) (l_e_st_eq_landau_n_rt_rp_rpofpd k q) i))))). Time Defined. (* constant 3251 *) Definition l_e_st_eq_landau_n_rt_rp_eqpdirp : (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd h), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (q:l_e_st_eq_landau_n_rt_rp_posd k), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd h p) (l_e_st_eq_landau_n_rt_rp_rpofpd k q)), l_e_st_eq_landau_n_rt_rp_eq h k))))). exact (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd h) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd k) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd h p) (l_e_st_eq_landau_n_rt_rp_rpofpd k q)) => l_e_st_eq_landau_n_rt_rp_tr3eq h (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd h p)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd k q)) k (l_e_st_eq_landau_n_rt_rp_eqpdrp1 h p) (l_e_st_eq_landau_n_rt_rp_iv1d_t9 h p k q i) (l_e_st_eq_landau_n_rt_rp_eqpdrp2 k q)))))). Time Defined. (* constant 3252 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t10 : (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd h), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (o:l_e_st_eq_landau_n_rt_rp_negd k), (forall (e:l_e_st_eq_landau_n_rt_rp_eq h k), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd h n)) (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd k o))))))). exact (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd h) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd k) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq h k) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd h n)) h k (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd k o)) (l_e_st_eq_landau_n_rt_rp_eqndrp2 h n) e (l_e_st_eq_landau_n_rt_rp_eqndrp1 k o)))))). Time Defined. (* constant 3253 *) Definition l_e_st_eq_landau_n_rt_rp_eqnderp : (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd h), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (o:l_e_st_eq_landau_n_rt_rp_negd k), (forall (e:l_e_st_eq_landau_n_rt_rp_eq h k), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnd h n) (l_e_st_eq_landau_n_rt_rp_rpofnd k o)))))). exact (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd h) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd k) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq h k) => l_e_st_eq_landau_n_rt_rp_isrpind (l_e_st_eq_landau_n_rt_rp_rpofnd h n) (l_e_st_eq_landau_n_rt_rp_rpofnd k o) (l_e_st_eq_landau_n_rt_rp_iv1d_t10 h n k o e)))))). Time Defined. (* constant 3254 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t11 : (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd h), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (o:l_e_st_eq_landau_n_rt_rp_negd k), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnd h n) (l_e_st_eq_landau_n_rt_rp_rpofnd k o)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd h n)) (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd k o))))))). exact (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd h) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd k) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnd h n) (l_e_st_eq_landau_n_rt_rp_rpofnd k o)) => l_e_st_eq_landau_n_rt_rp_isrpend (l_e_st_eq_landau_n_rt_rp_rpofnd h n) (l_e_st_eq_landau_n_rt_rp_rpofnd k o) i))))). Time Defined. (* constant 3255 *) Definition l_e_st_eq_landau_n_rt_rp_eqndirp : (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd h), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (o:l_e_st_eq_landau_n_rt_rp_negd k), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnd h n) (l_e_st_eq_landau_n_rt_rp_rpofnd k o)), l_e_st_eq_landau_n_rt_rp_eq h k))))). exact (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd h) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd k) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnd h n) (l_e_st_eq_landau_n_rt_rp_rpofnd k o)) => l_e_st_eq_landau_n_rt_rp_tr3eq h (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd h n)) (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd k o)) k (l_e_st_eq_landau_n_rt_rp_eqndrp1 h n) (l_e_st_eq_landau_n_rt_rp_iv1d_t11 h n k o i) (l_e_st_eq_landau_n_rt_rp_eqndrp2 k o)))))). Time Defined. (* constant 3256 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t12 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_posdirp r)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_eqpdrp1 (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_posdirp r)). Time Defined. (* constant 3257 *) Definition l_e_st_eq_landau_n_rt_rp_isrppd1 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_posdirp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isrpipd r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_posdirp r)) (l_e_st_eq_landau_n_rt_rp_iv1d_t12 r)). Time Defined. (* constant 3258 *) Definition l_e_st_eq_landau_n_rt_rp_isrppd2 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_posdirp r)) r). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_posdirp r)) (l_e_st_eq_landau_n_rt_rp_isrppd1 r)). Time Defined. (* constant 3259 *) Definition l_e_st_eq_landau_n_rt_rp_iv1d_t13 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_rpofnd (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_negdirp r)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_eqndrp1 (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_negdirp r)). Time Defined. (* constant 3260 *) Definition l_e_st_eq_landau_n_rt_rp_isrpnd1 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is r (l_e_st_eq_landau_n_rt_rp_rpofnd (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_negdirp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_isrpind r (l_e_st_eq_landau_n_rt_rp_rpofnd (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_negdirp r)) (l_e_st_eq_landau_n_rt_rp_iv1d_t13 r)). Time Defined. (* constant 3261 *) Definition l_e_st_eq_landau_n_rt_rp_isrpnd2 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnd (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_negdirp r)) r). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut r (l_e_st_eq_landau_n_rt_rp_rpofnd (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_negdirp r)) (l_e_st_eq_landau_n_rt_rp_isrpnd1 r)). Time Defined. (* constant 3262 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad1 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 r) (l_e_st_eq_landau_n_rt_rp_pl a2 r))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_eqi12 a1 a2 (l_e_st_eq_landau_n_rt_rp_pl a1 r) (l_e_st_eq_landau_n_rt_rp_pl a2 r) (l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl a1 (l_e_st_eq_landau_n_rt_rp_pl a2 r)) (l_e_st_eq_landau_n_rt_rp_pl a1 (l_e_st_eq_landau_n_rt_rp_pl r a2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl a1 r) a2) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl a2 r) (l_e_st_eq_landau_n_rt_rp_pl r a2) a1 (l_e_st_eq_landau_n_rt_rp_compl a2 r)) (l_e_st_eq_landau_n_rt_rp_asspl2 a1 r a2))))). Time Defined. (* constant 3263 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad2 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 r) (l_e_st_eq_landau_n_rt_rp_pl a2 r)) (l_e_st_eq_landau_n_rt_rp_df a1 a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 r) (l_e_st_eq_landau_n_rt_rp_pl a2 r)) (l_e_st_eq_landau_n_rt_rp_lemmad1 a1 a2 r)))). Time Defined. (* constant 3264 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_treq a (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r)) (l_e_st_eq_landau_n_rt_rp_refeq1 a (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_isdf a)) (l_e_st_eq_landau_n_rt_rp_lemmad1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) r))). Time Defined. (* constant 3265 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r)) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_symeq a (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r)) (l_e_st_eq_landau_n_rt_rp_lemmad3 a r))). Time Defined. (* constant 3266 *) Definition l_e_st_eq_landau_n_rt_rp_absd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_dif). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_ite (l_e_st_eq_landau_n_rt_rp_negd a) l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) a). Time Defined. (* constant 3267 *) Definition l_e_st_eq_landau_n_rt_rp_absnd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_itet (l_e_st_eq_landau_n_rt_rp_negd a) l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) a n))). Time Defined. (* constant 3268 *) Definition l_e_st_eq_landau_n_rt_rp_absnnd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_itef (l_e_st_eq_landau_n_rt_rp_negd a) l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) a n))). Time Defined. (* constant 3269 *) Definition l_e_st_eq_landau_n_rt_rp_absdeql : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less a1 a2), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_df a2 a1)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less a1 a2) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_df a2 a1) (l_e_st_eq_landau_n_rt_rp_absnd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_negdi a1 a2 l)) (l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2 a1 (l_e_st_eq_landau_n_rt_rp_stdis a1 a2) (l_e_st_eq_landau_n_rt_rp_stmis a1 a2))))). Time Defined. (* constant 3270 *) Definition l_e_st_eq_landau_n_rt_rp_absdeqm : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_moreis a1 a2), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_df a1 a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreis a1 a2) => l_e_st_eq_landau_n_rt_rp_absnnd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_less a1 a2) (l_e_st_eq_landau_n_rt_rp_satz123c a1 a2 m) (fun (t:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_df a1 a2)) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a1 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2 (l_e_st_eq_landau_n_rt_rp_stmis a1 a2) (l_e_st_eq_landau_n_rt_rp_stdis a1 a2) t))))). Time Defined. (* constant 3271 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_symeq a b e) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))))). Time Defined. (* constant 3272 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_absnd a n) (l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_iv2d_t1 a b e n)) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_absnd b (l_e_st_eq_landau_n_rt_rp_eqnegd a b e n))))))). Time Defined. (* constant 3273 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_absd a) a b (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_absnnd a n) e (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd b) b (l_e_st_eq_landau_n_rt_rp_absnnd b (l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_negd a) n (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_eqnegd b a (l_e_st_eq_landau_n_rt_rp_symeq a b e) t)))))))). Time Defined. (* constant 3274 *) Definition l_e_st_eq_landau_n_rt_rp_eqabsd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_iv2d_t2 a b e t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => l_e_st_eq_landau_n_rt_rp_iv2d_t3 a b e t)))). Time Defined. (* constant 3275 *) Definition l_e_st_eq_landau_n_rt_rp_satzd166a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_eqposd a (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a p))) p)). Time Defined. (* constant 3276 *) Definition l_e_st_eq_landau_n_rt_rp_2d166_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_posdi (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) n))). Time Defined. (* constant 3277 *) Definition l_e_st_eq_landau_n_rt_rp_satzd166b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_absnd a n)) (l_e_st_eq_landau_n_rt_rp_2d166_t1 a n))). Time Defined. (* constant 3278 *) Definition l_e_st_eq_landau_n_rt_rp_satzd166c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)), l_e_st_eq_landau_n_rt_rp_eq a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) => l_e_st_eq_landau_n_rt_rp_tr3eq a (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b) b (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a p))) e (l_e_st_eq_landau_n_rt_rp_absnnd b (l_e_st_eq_landau_n_rt_rp_pnotnd b q))))))). Time Defined. (* constant 3279 *) Definition l_e_st_eq_landau_n_rt_rp_2d166_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (o:l_e_st_eq_landau_n_rt_rp_negd b), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd b) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_absnd a n)) e (l_e_st_eq_landau_n_rt_rp_absnd b o)))))). Time Defined. (* constant 3280 *) Definition l_e_st_eq_landau_n_rt_rp_satzd166d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (o:l_e_st_eq_landau_n_rt_rp_negd b), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)), l_e_st_eq_landau_n_rt_rp_eq a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd b) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_eqe12 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d166_t2 a b n o e))) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))))))). Time Defined. (* constant 3281 *) Definition l_e_st_eq_landau_n_rt_rp_satzd166e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_rappd a (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_satzd166a a t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a)) n) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_satzd166b a t))). Time Defined. (* constant 3282 *) Definition l_e_st_eq_landau_n_rt_rp_satzd166f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_eqzero a (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_0notnd a z))) z)). Time Defined. (* constant 3283 *) Definition l_e_st_eq_landau_n_rt_rp_mored : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3284 *) Definition l_e_st_eq_landau_n_rt_rp_moredi12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_pl b1 a2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_12issmsd a1 a2 b1 b2) (l_e_st_eq_landau_n_rt_rp_12issmsd b1 b2 a1 a2) m))))). Time Defined. (* constant 3285 *) Definition l_e_st_eq_landau_n_rt_rp_moredi1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a))), l_e_st_eq_landau_n_rt_rp_mored a (l_e_st_eq_landau_n_rt_rp_df r1 r2))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a))) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2))) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_sm2issmsd a r1 r2) (l_e_st_eq_landau_n_rt_rp_1sdissmsd a r1 r2) m)))). Time Defined. (* constant 3286 *) Definition l_e_st_eq_landau_n_rt_rp_moredi2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2)), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_df r1 r2) a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2)) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2))) (l_e_st_eq_landau_n_rt_rp_1sdissmsd a r1 r2) (l_e_st_eq_landau_n_rt_rp_sm2issmsd a r1 r2) m)))). Time Defined. (* constant 3287 *) Definition l_e_st_eq_landau_n_rt_rp_morede12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_pl b1 a2) (l_e_st_eq_landau_n_rt_rp_smsdis12 a1 a2 b1 b2) (l_e_st_eq_landau_n_rt_rp_smsdis12 b1 b2 a1 a2) m))))). Time Defined. (* constant 3288 *) Definition l_e_st_eq_landau_n_rt_rp_lessd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3289 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_lessd b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) m))). Time Defined. (* constant 3290 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_mored b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) l))). Time Defined. (* constant 3291 *) Definition l_e_st_eq_landau_n_rt_rp_lessdi12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)) => l_e_st_eq_landau_n_rt_rp_lemmad5 (l_e_st_eq_landau_n_rt_rp_df b1 b2) (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_moredi12 b1 b2 a1 a2 (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2) l))))))). Time Defined. (* constant 3292 *) Definition l_e_st_eq_landau_n_rt_rp_lessdi1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a))), l_e_st_eq_landau_n_rt_rp_lessd a (l_e_st_eq_landau_n_rt_rp_df r1 r2))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a))) => l_e_st_eq_landau_n_rt_rp_lemmad5 (l_e_st_eq_landau_n_rt_rp_df r1 r2) a (l_e_st_eq_landau_n_rt_rp_moredi2 a r1 r2 (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) l)))))). Time Defined. (* constant 3293 *) Definition l_e_st_eq_landau_n_rt_rp_lessdi2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2)), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_df r1 r2) a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2)) => l_e_st_eq_landau_n_rt_rp_lemmad5 a (l_e_st_eq_landau_n_rt_rp_df r1 r2) (l_e_st_eq_landau_n_rt_rp_moredi1 a r1 r2 (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) l)))))). Time Defined. (* constant 3294 *) Definition l_e_st_eq_landau_n_rt_rp_lessde12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl b1 a2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_pl b1 a2) (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_morede12 b1 b2 a1 a2 (l_e_st_eq_landau_n_rt_rp_lemmad6 (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2) l))))))). Time Defined. (* constant 3295 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_orec3 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satz123 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3296 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_or3 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satz123a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3297 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_ec3 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3298 *) Definition l_e_st_eq_landau_n_rt_rp_1d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_stm d)))). Time Defined. (* constant 3299 *) Definition l_e_st_eq_landau_n_rt_rp_2d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_cut)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_std d)))). Time Defined. (* constant 3300 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a c), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d a b c d)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a c) => l_e_tr4is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d a b c d)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2d a b c d))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d a b c d) (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2d a b c d)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_symeq a b e) f) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2b a b)))))))))). Time Defined. (* constant 3301 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a c), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d a b c d)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a c) => l_e_st_eq_landau_n_rt_rp_ismore2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d a b c d)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d a b c d)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_iv2d_t4 a b c d e f m) (l_e_st_eq_landau_n_rt_rp_satz135d (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d a b c d)) m)))))))). Time Defined. (* constant 3302 *) Definition l_e_st_eq_landau_n_rt_rp_eqmored12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a c), l_e_st_eq_landau_n_rt_rp_mored b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a c) => l_e_st_eq_landau_n_rt_rp_satz136a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2d a b c d)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1d a b c d) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_iv2d_t5 a b c d e f m)))))))). Time Defined. (* constant 3303 *) Definition l_e_st_eq_landau_n_rt_rp_eqlessd12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a c), l_e_st_eq_landau_n_rt_rp_lessd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a c) => l_e_st_eq_landau_n_rt_rp_lemmad5 d b (l_e_st_eq_landau_n_rt_rp_eqmored12 c d a b f e (l_e_st_eq_landau_n_rt_rp_lemmad6 a c l))))))))). Time Defined. (* constant 3304 *) Definition l_e_st_eq_landau_n_rt_rp_eqmored1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a c), l_e_st_eq_landau_n_rt_rp_mored b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a c) => l_e_st_eq_landau_n_rt_rp_eqmored12 a b c c e (l_e_st_eq_landau_n_rt_rp_refeq c) m))))). Time Defined. (* constant 3305 *) Definition l_e_st_eq_landau_n_rt_rp_eqmored2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (m:l_e_st_eq_landau_n_rt_rp_mored c a), l_e_st_eq_landau_n_rt_rp_mored c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored c a) => l_e_st_eq_landau_n_rt_rp_eqmored12 c c a b (l_e_st_eq_landau_n_rt_rp_refeq c) e m))))). Time Defined. (* constant 3306 *) Definition l_e_st_eq_landau_n_rt_rp_eqlessd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a c), l_e_st_eq_landau_n_rt_rp_lessd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a c) => l_e_st_eq_landau_n_rt_rp_eqlessd12 a b c c e (l_e_st_eq_landau_n_rt_rp_refeq c) l))))). Time Defined. (* constant 3307 *) Definition l_e_st_eq_landau_n_rt_rp_eqlessd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd c a), l_e_st_eq_landau_n_rt_rp_lessd c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd c a) => l_e_st_eq_landau_n_rt_rp_eqlessd12 c c a b (l_e_st_eq_landau_n_rt_rp_refeq c) e l))))). Time Defined. (* constant 3308 *) Definition l_e_st_eq_landau_n_rt_rp_moreq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_or (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b))). Time Defined. (* constant 3309 *) Definition l_e_st_eq_landau_n_rt_rp_lesseq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_or (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_eq a b))). Time Defined. (* constant 3310 *) Definition l_e_st_eq_landau_n_rt_rp_satzd168a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), l_e_st_eq_landau_n_rt_rp_lesseq b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_lessd b a) (l_e_st_eq_landau_n_rt_rp_eq b a) m (fun (t:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_lemmad5 a b t) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_symeq a b t)))). Time Defined. (* constant 3311 *) Definition l_e_st_eq_landau_n_rt_rp_satzd168b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), l_e_st_eq_landau_n_rt_rp_moreq b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored b a) (l_e_st_eq_landau_n_rt_rp_eq b a) l (fun (t:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_lemmad6 a b t) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_symeq a b t)))). Time Defined. (* constant 3312 *) Definition l_e_st_eq_landau_n_rt_rp_eqmoreq1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a c), l_e_st_eq_landau_n_rt_rp_moreq b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a c) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_mored a c) (l_e_st_eq_landau_n_rt_rp_eq a c) (l_e_st_eq_landau_n_rt_rp_mored b c) (l_e_st_eq_landau_n_rt_rp_eq b c) m (fun (t:l_e_st_eq_landau_n_rt_rp_mored a c) => l_e_st_eq_landau_n_rt_rp_eqmored1 a b c e t) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a c) => l_e_st_eq_landau_n_rt_rp_treq1 b c a e t)))))). Time Defined. (* constant 3313 *) Definition l_e_st_eq_landau_n_rt_rp_eqmoreq2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq c a), l_e_st_eq_landau_n_rt_rp_moreq c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq c a) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_mored c a) (l_e_st_eq_landau_n_rt_rp_eq c a) (l_e_st_eq_landau_n_rt_rp_mored c b) (l_e_st_eq_landau_n_rt_rp_eq c b) m (fun (t:l_e_st_eq_landau_n_rt_rp_mored c a) => l_e_st_eq_landau_n_rt_rp_eqmored2 a b c e t) (fun (t:l_e_st_eq_landau_n_rt_rp_eq c a) => l_e_st_eq_landau_n_rt_rp_treq c a b t e)))))). Time Defined. (* constant 3314 *) Definition l_e_st_eq_landau_n_rt_rp_eqlesseq1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a c), l_e_st_eq_landau_n_rt_rp_lesseq b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a c) => l_e_st_eq_landau_n_rt_rp_satzd168a c b (l_e_st_eq_landau_n_rt_rp_eqmoreq2 a b c e (l_e_st_eq_landau_n_rt_rp_satzd168b a c l))))))). Time Defined. (* constant 3315 *) Definition l_e_st_eq_landau_n_rt_rp_eqlesseq2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq c a), l_e_st_eq_landau_n_rt_rp_lesseq c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq c a) => l_e_st_eq_landau_n_rt_rp_satzd168a b c (l_e_st_eq_landau_n_rt_rp_eqmoreq1 a b c e (l_e_st_eq_landau_n_rt_rp_satzd168b c a l))))))). Time Defined. (* constant 3316 *) Definition l_e_st_eq_landau_n_rt_rp_eqmoreq12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a c), l_e_st_eq_landau_n_rt_rp_moreq b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a c) => l_e_st_eq_landau_n_rt_rp_eqmoreq1 a b d e (l_e_st_eq_landau_n_rt_rp_eqmoreq2 c d a f m)))))))). Time Defined. (* constant 3317 *) Definition l_e_st_eq_landau_n_rt_rp_eqlesseq12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a c), l_e_st_eq_landau_n_rt_rp_lesseq b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a c) => l_e_st_eq_landau_n_rt_rp_eqlesseq1 a b d e (l_e_st_eq_landau_n_rt_rp_eqlesseq2 c d a f l)))))))). Time Defined. (* constant 3318 *) Definition l_e_st_eq_landau_n_rt_rp_moreqi1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_moreq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_ori1 (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b) m))). Time Defined. (* constant 3319 *) Definition l_e_st_eq_landau_n_rt_rp_lesseqi1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_lesseq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_ori1 (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_eq a b) l))). Time Defined. (* constant 3320 *) Definition l_e_st_eq_landau_n_rt_rp_moreqi2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_moreq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_ori2 (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b) e))). Time Defined. (* constant 3321 *) Definition l_e_st_eq_landau_n_rt_rp_lesseqi2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_lesseq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_ori2 (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_eq a b) e))). Time Defined. (* constant 3322 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), l_not (l_e_st_eq_landau_n_rt_rp_lessd a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => l_ec3_th7 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_satzd167b a b) (l_comor (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b) m)))). Time Defined. (* constant 3323 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), l_not (l_e_st_eq_landau_n_rt_rp_mored a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => l_ec3_th9 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_satzd167b a b) l))). Time Defined. (* constant 3324 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_mored a b)), l_e_st_eq_landau_n_rt_rp_lesseq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_mored a b)) => l_or3_th2 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_satzd167a a b) n))). Time Defined. (* constant 3325 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_lessd a b)), l_e_st_eq_landau_n_rt_rp_moreq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_lessd a b)) => l_comor (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_or3_th3 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_satzd167a a b) n)))). Time Defined. (* constant 3326 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167g : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_not (l_e_st_eq_landau_n_rt_rp_lesseq a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_lesseq a b) (l_not (l_e_st_eq_landau_n_rt_rp_mored a b)) (l_weli (l_e_st_eq_landau_n_rt_rp_mored a b) m) (fun (t:l_e_st_eq_landau_n_rt_rp_lesseq a b) => l_e_st_eq_landau_n_rt_rp_satzd167d a b t)))). Time Defined. (* constant 3327 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167h : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_not (l_e_st_eq_landau_n_rt_rp_moreq a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_moreq a b) (l_not (l_e_st_eq_landau_n_rt_rp_lessd a b)) (l_weli (l_e_st_eq_landau_n_rt_rp_lessd a b) l) (fun (t:l_e_st_eq_landau_n_rt_rp_moreq a b) => l_e_st_eq_landau_n_rt_rp_satzd167c a b t)))). Time Defined. (* constant 3328 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167j : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_moreq a b)), l_e_st_eq_landau_n_rt_rp_lessd a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_moreq a b)) => l_or3e3 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_satzd167a a b) (l_or_th5 (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b) n) (l_or_th4 (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b) n)))). Time Defined. (* constant 3329 *) Definition l_e_st_eq_landau_n_rt_rp_satzd167k : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_lesseq a b)), l_e_st_eq_landau_n_rt_rp_mored a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_lesseq a b)) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_satzd167a a b) (l_or_th4 (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_eq a b) n) (l_or_th5 (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_eq a b) n)))). Time Defined. (* constant 3330 *) Definition l_e_st_eq_landau_n_rt_rp_satzd169a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_mored a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a) z) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_satz135a (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b) p))))). Time Defined. (* constant 3331 *) Definition l_e_st_eq_landau_n_rt_rp_satzd169b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_posd a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_satz136d (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a) z) m))))). Time Defined. (* constant 3332 *) Definition l_e_st_eq_landau_n_rt_rp_satzd169c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_lessd a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a) z) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_satz135c (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b) n))))). Time Defined. (* constant 3333 *) Definition l_e_st_eq_landau_n_rt_rp_satzd169d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_negd a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_satz136f (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a) z) l))))). Time Defined. (* constant 3334 *) Definition l_e_st_eq_landau_n_rt_rp_2d170_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_absd a) b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_moreqi1 (l_e_st_eq_landau_n_rt_rp_absd a) b (l_e_st_eq_landau_n_rt_rp_satzd169a (l_e_st_eq_landau_n_rt_rp_absd a) b z (l_e_st_eq_landau_n_rt_rp_satzd166a a p)))))). Time Defined. (* constant 3335 *) Definition l_e_st_eq_landau_n_rt_rp_2d170_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), (forall (y:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_absd a) b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => (fun (y:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_moreqi2 (l_e_st_eq_landau_n_rt_rp_absd a) b (l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_absd a) a b (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_0notnd a y)) (l_e_st_eq_landau_n_rt_rp_zeroeq a b y z)))))). Time Defined. (* constant 3336 *) Definition l_e_st_eq_landau_n_rt_rp_2d170_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_absd a) b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_moreqi1 (l_e_st_eq_landau_n_rt_rp_absd a) b (l_e_st_eq_landau_n_rt_rp_satzd169a (l_e_st_eq_landau_n_rt_rp_absd a) b z (l_e_st_eq_landau_n_rt_rp_satzd166b a n)))))). Time Defined. (* constant 3337 *) Definition l_e_st_eq_landau_n_rt_rp_satzd170 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_absd a) b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_rappd a (l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_absd a) b) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_2d170_t1 a b z t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_2d170_t2 a b z t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_2d170_t3 a b z t)))). Time Defined. (* constant 3338 *) Definition l_e_st_eq_landau_n_rt_rp_2d171_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lessd b c), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessd b c) => l_e_st_eq_landau_n_rt_rp_satz137a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b)) l k))))). Time Defined. (* constant 3339 *) Definition l_e_st_eq_landau_n_rt_rp_2d171_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lessd b c), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessd b c) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_2d171_t1 a b c l k)))))). Time Defined. (* constant 3340 *) Definition l_e_st_eq_landau_n_rt_rp_satzd171 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lessd b c), l_e_st_eq_landau_n_rt_rp_lessd a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessd b c) => l_e_st_eq_landau_n_rt_rp_satz136c (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_2d171_t2 a b c l k)))))). Time Defined. (* constant 3341 *) Definition l_e_st_eq_landau_n_rt_rp_trlessd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lessd b c), l_e_st_eq_landau_n_rt_rp_lessd a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessd b c) => l_e_st_eq_landau_n_rt_rp_satzd171 a b c l k))))). Time Defined. (* constant 3342 *) Definition l_e_st_eq_landau_n_rt_rp_trmored : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (n:l_e_st_eq_landau_n_rt_rp_mored b c), l_e_st_eq_landau_n_rt_rp_mored a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_mored b c) => l_e_st_eq_landau_n_rt_rp_lemmad6 c a (l_e_st_eq_landau_n_rt_rp_trlessd c b a (l_e_st_eq_landau_n_rt_rp_lemmad5 b c n) (l_e_st_eq_landau_n_rt_rp_lemmad5 a b m))))))). Time Defined. (* constant 3343 *) Definition l_e_st_eq_landau_n_rt_rp_satzd172a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lessd b c), l_e_st_eq_landau_n_rt_rp_lessd a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessd b c) => l_orapp (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_lessd a c) l (fun (t:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_trlessd a b c t k) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqlessd1 b a c (l_e_st_eq_landau_n_rt_rp_symeq a b t) k)))))). Time Defined. (* constant 3344 *) Definition l_e_st_eq_landau_n_rt_rp_satzd172b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lesseq b c), l_e_st_eq_landau_n_rt_rp_lessd a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lesseq b c) => l_orapp (l_e_st_eq_landau_n_rt_rp_lessd b c) (l_e_st_eq_landau_n_rt_rp_eq b c) (l_e_st_eq_landau_n_rt_rp_lessd a c) k (fun (t:l_e_st_eq_landau_n_rt_rp_lessd b c) => l_e_st_eq_landau_n_rt_rp_trlessd a b c l t) (fun (t:l_e_st_eq_landau_n_rt_rp_eq b c) => l_e_st_eq_landau_n_rt_rp_eqlessd2 b c a t l)))))). Time Defined. (* constant 3345 *) Definition l_e_st_eq_landau_n_rt_rp_satzd172c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_mored b c), l_e_st_eq_landau_n_rt_rp_mored a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_mored b c) => l_e_st_eq_landau_n_rt_rp_lemmad6 c a (l_e_st_eq_landau_n_rt_rp_satzd172b c b a (l_e_st_eq_landau_n_rt_rp_lemmad5 b c n) (l_e_st_eq_landau_n_rt_rp_satzd168a a b m))))))). Time Defined. (* constant 3346 *) Definition l_e_st_eq_landau_n_rt_rp_satzd172d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (n:l_e_st_eq_landau_n_rt_rp_moreq b c), l_e_st_eq_landau_n_rt_rp_mored a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreq b c) => l_e_st_eq_landau_n_rt_rp_lemmad6 c a (l_e_st_eq_landau_n_rt_rp_satzd172a c b a (l_e_st_eq_landau_n_rt_rp_satzd168a b c n) (l_e_st_eq_landau_n_rt_rp_lemmad5 a b m))))))). Time Defined. (* constant 3347 *) Definition l_e_st_eq_landau_n_rt_rp_2d173_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lesseq b c), (forall (j:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_lesseq a c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lesseq b c) => (fun (j:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_lesseqi1 a c (l_e_st_eq_landau_n_rt_rp_satzd172b a b c j k))))))). Time Defined. (* constant 3348 *) Definition l_e_st_eq_landau_n_rt_rp_2d173_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lesseq b c), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_lesseq a c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lesseq b c) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqlesseq1 b a c (l_e_st_eq_landau_n_rt_rp_symeq a b e) k)))))). Time Defined. (* constant 3349 *) Definition l_e_st_eq_landau_n_rt_rp_satzd173 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lesseq b c), l_e_st_eq_landau_n_rt_rp_lesseq a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lesseq b c) => l_orapp (l_e_st_eq_landau_n_rt_rp_lessd a b) (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_lesseq a c) l (fun (t:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_2d173_t1 a b c l k t) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_2d173_t2 a b c l k t)))))). Time Defined. (* constant 3350 *) Definition l_e_st_eq_landau_n_rt_rp_trlesseq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lesseq b c), l_e_st_eq_landau_n_rt_rp_lesseq a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lesseq b c) => l_e_st_eq_landau_n_rt_rp_satzd173 a b c l k))))). Time Defined. (* constant 3351 *) Definition l_e_st_eq_landau_n_rt_rp_trmoreq : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_moreq b c), l_e_st_eq_landau_n_rt_rp_moreq a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreq b c) => l_e_st_eq_landau_n_rt_rp_satzd168b c a (l_e_st_eq_landau_n_rt_rp_trlesseq c b a (l_e_st_eq_landau_n_rt_rp_satzd168a b c n) (l_e_st_eq_landau_n_rt_rp_satzd168a a b m))))))). Time Defined. (* constant 3352 *) Definition l_e_st_eq_landau_n_rt_rp_ratd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (forall (t:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_satzd166e a t)))). Time Defined. (* constant 3353 *) Definition l_e_st_eq_landau_n_rt_rp_irratd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_not (l_e_st_eq_landau_n_rt_rp_ratd a)). Time Defined. (* constant 3354 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (r:l_e_st_eq_landau_n_rt_rp_ratd a), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), l_not (l_e_st_eq_landau_n_rt_rp_zero a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (r:l_e_st_eq_landau_n_rt_rp_ratd a) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_zero b) n (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_eqzero a b e t)))))). Time Defined. (* constant 3355 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (r:l_e_st_eq_landau_n_rt_rp_ratd a), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_satzd166e a (l_e_st_eq_landau_n_rt_rp_iv2d_t6 a b e r n))) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_satzd166e b n))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (r:l_e_st_eq_landau_n_rt_rp_ratd a) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_e_st_eq_landau_n_rt_rp_eqpderp (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_satzd166e a (l_e_st_eq_landau_n_rt_rp_iv2d_t6 a b e r n)) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_satzd166e b n) (l_e_st_eq_landau_n_rt_rp_eqabsd a b e)))))). Time Defined. (* constant 3356 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t8 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (r:l_e_st_eq_landau_n_rt_rp_ratd a), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_satzd166e b n))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (r:l_e_st_eq_landau_n_rt_rp_ratd a) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp t) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_satzd166e a (l_e_st_eq_landau_n_rt_rp_iv2d_t6 a b e r n))) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_satzd166e b n)) (r (l_e_st_eq_landau_n_rt_rp_iv2d_t6 a b e r n)) (l_e_st_eq_landau_n_rt_rp_iv2d_t7 a b e r n)))))). Time Defined. (* constant 3357 *) Definition l_e_st_eq_landau_n_rt_rp_eqratd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (r:l_e_st_eq_landau_n_rt_rp_ratd a), l_e_st_eq_landau_n_rt_rp_ratd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (r:l_e_st_eq_landau_n_rt_rp_ratd a) => (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_e_st_eq_landau_n_rt_rp_iv2d_t8 a b e r t))))). Time Defined. (* constant 3358 *) Definition l_e_st_eq_landau_n_rt_rp_eqirratd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (i:l_e_st_eq_landau_n_rt_rp_irratd a), l_e_st_eq_landau_n_rt_rp_irratd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratd a) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_ratd b) (l_e_st_eq_landau_n_rt_rp_ratd a) i (fun (t:l_e_st_eq_landau_n_rt_rp_ratd b) => l_e_st_eq_landau_n_rt_rp_eqratd b a (l_e_st_eq_landau_n_rt_rp_symeq a b e) t))))). Time Defined. (* constant 3359 *) Definition l_e_st_eq_landau_n_rt_rp_ratdi0 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_ratd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_r_imp_th2 (l_not (l_e_st_eq_landau_n_rt_rp_zero a)) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_satzd166e a t))) (l_weli (l_e_st_eq_landau_n_rt_rp_zero a) z))). Time Defined. (* constant 3360 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t9 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) r) (l_e_st_eq_landau_n_rt_rp_rpofrt y0))))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) r) (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) r) j)))))). Time Defined. (* constant 3361 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t10 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt x0))))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_e_st_eq_landau_n_rt_rp_ismore1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofrt x0) r) (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_iv2d_t9 r i x0 s y0 j) (l_e_st_eq_landau_n_rt_rp_satz133 (l_e_st_eq_landau_n_rt_rp_rpofrt x0) r))))))). Time Defined. (* constant 3362 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t11 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_more y0 x0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_e_st_eq_landau_n_rt_rp_satz154d y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t10 r i x0 s y0 j))))))). Time Defined. (* constant 3363 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t12 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j))) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_satz154a y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_e_st_eq_landau_n_rt_rp_satz155b y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j))))))). Time Defined. (* constant 3364 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t13 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_rp_is r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_satz154a y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_e_st_eq_landau_n_rt_rp_satz140g (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) r (l_e_st_eq_landau_n_rt_rp_satz154a y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j)) (l_e_st_eq_landau_n_rt_rp_iv2d_t9 r i x0 s y0 j))))))). Time Defined. (* constant 3365 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t14 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_rp_is r (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_e_tris2 l_e_st_eq_landau_n_rt_cut r (l_e_st_eq_landau_n_rt_rp_rpofrt (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j))) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_rpofrt y0) (l_e_st_eq_landau_n_rt_rp_rpofrt x0) (l_e_st_eq_landau_n_rt_rp_satz154a y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j))) (l_e_st_eq_landau_n_rt_rp_iv2d_t13 r i x0 s y0 j) (l_e_st_eq_landau_n_rt_rp_iv2d_t12 r i x0 s y0 j))))))). Time Defined. (* constant 3366 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t15 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)), l_e_st_eq_landau_n_rt_rp_ratrp r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (j:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt y0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_is r (l_e_st_eq_landau_n_rt_rp_rpofrt x)) (l_e_st_eq_landau_n_rt_mn y0 x0 (l_e_st_eq_landau_n_rt_rp_iv2d_t11 r i x0 s y0 j)) (l_e_st_eq_landau_n_rt_rp_iv2d_t14 r i x0 s y0 j))))))). Time Defined. (* constant 3367 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t16 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))), l_con)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (s:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => l_someapp l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) s l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0)) (l_e_st_eq_landau_n_rt_rp_rpofrt x)) => i (l_e_st_eq_landau_n_rt_rp_iv2d_t15 r i x0 s x t))))))). Time Defined. (* constant 3368 *) Definition l_e_st_eq_landau_n_rt_rp_remark1 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_rp_irratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_rpofrt x0))) => l_e_st_eq_landau_n_rt_rp_iv2d_t16 r i x0 t)))). Time Defined. (* constant 3369 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_rp : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_dif). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pdofrp r). Time Defined. (* constant 3370 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_rn : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_dif). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ndofrp r). Time Defined. (* constant 3371 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t17 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_iv2d_rp r)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_posdirp r). Time Defined. (* constant 3372 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t18 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_iv2d_rp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pnot0d (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t17 r)). Time Defined. (* constant 3373 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t19 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_iv2d_rn r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_nnot0d (l_e_st_eq_landau_n_rt_rp_iv2d_rn r) (l_e_st_eq_landau_n_rt_rp_negdirp r)). Time Defined. (* constant 3374 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t20 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_iv2d_rp r))), l_e_st_eq_landau_n_rt_rp_is r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_iv2d_rp r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) n)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_iv2d_rp r))) => l_e_tris2 l_e_st_eq_landau_n_rt_cut r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_iv2d_rp r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) n)) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t17 r)) (l_e_st_eq_landau_n_rt_rp_isrppd1 r) (l_e_st_eq_landau_n_rt_rp_eqpderp (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_iv2d_rp r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) n) (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t17 r) (l_e_st_eq_landau_n_rt_rp_absnnd (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) (l_e_st_eq_landau_n_rt_rp_pnotnd (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t17 r)))))). Time Defined. (* constant 3375 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t21 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_iv2d_rn r)) (l_e_st_eq_landau_n_rt_rp_iv2d_rp r)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_iv2d_rn r)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_iv2d_rn r)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_iv2d_rn r))) (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) (l_e_st_eq_landau_n_rt_rp_absnd (l_e_st_eq_landau_n_rt_rp_iv2d_rn r) (l_e_st_eq_landau_n_rt_rp_negdirp r)) (l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_iv2d_rn r)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_iv2d_rn r)) (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_stdis l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_stmis l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp)))). Time Defined. (* constant 3376 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t22 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_iv2d_rn r))), l_e_st_eq_landau_n_rt_rp_is r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_iv2d_rn r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_iv2d_rn r) n)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_iv2d_rn r))) => l_e_tris2 l_e_st_eq_landau_n_rt_cut r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_iv2d_rn r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_iv2d_rn r) n)) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t17 r)) (l_e_st_eq_landau_n_rt_rp_isrppd1 r) (l_e_st_eq_landau_n_rt_rp_eqpderp (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_iv2d_rn r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_iv2d_rn r) n) (l_e_st_eq_landau_n_rt_rp_iv2d_rp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t17 r) (l_e_st_eq_landau_n_rt_rp_iv2d_t21 r)))). Time Defined. (* constant 3377 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t23 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is r s), (forall (rr:l_e_st_eq_landau_n_rt_rp_ratrp r), l_e_st_eq_landau_n_rt_rp_ratrp s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is r s) => (fun (rr:l_e_st_eq_landau_n_rt_rp_ratrp r) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp x) r s rr i)))). Time Defined. (* constant 3378 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t24 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is r s), (forall (rs:l_e_st_eq_landau_n_rt_rp_ratrp s), l_e_st_eq_landau_n_rt_rp_ratrp r)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is r s) => (fun (rs:l_e_st_eq_landau_n_rt_rp_ratrp s) => l_e_isp1 l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ratrp x) s r rs i)))). Time Defined. (* constant 3379 *) Definition l_e_st_eq_landau_n_rt_rp_remark2a : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (rr:l_e_st_eq_landau_n_rt_rp_ratrp r), l_e_st_eq_landau_n_rt_rp_ratd (l_e_st_eq_landau_n_rt_rp_pdofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (rr:l_e_st_eq_landau_n_rt_rp_ratrp r) => (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pdofrp r))) => l_e_st_eq_landau_n_rt_rp_iv2d_t23 r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_pdofrp r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_pdofrp r) t)) (l_e_st_eq_landau_n_rt_rp_iv2d_t20 r t) rr))). Time Defined. (* constant 3380 *) Definition l_e_st_eq_landau_n_rt_rp_remark2b : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (rr:l_e_st_eq_landau_n_rt_rp_ratrp r), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (rr:l_e_st_eq_landau_n_rt_rp_ratrp r) => l_e_st_eq_landau_n_rt_rp_iv2d_t17 r)). Time Defined. (* constant 3381 *) Definition l_e_st_eq_landau_n_rt_rp_remark3a : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (rr:l_e_st_eq_landau_n_rt_rp_ratrp r), l_e_st_eq_landau_n_rt_rp_ratd (l_e_st_eq_landau_n_rt_rp_ndofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (rr:l_e_st_eq_landau_n_rt_rp_ratrp r) => (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_ndofrp r))) => l_e_st_eq_landau_n_rt_rp_iv2d_t23 r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_ndofrp r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_ndofrp r) t)) (l_e_st_eq_landau_n_rt_rp_iv2d_t22 r t) rr))). Time Defined. (* constant 3382 *) Definition l_e_st_eq_landau_n_rt_rp_remark3b : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (rr:l_e_st_eq_landau_n_rt_rp_ratrp r), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_ndofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (rr:l_e_st_eq_landau_n_rt_rp_ratrp r) => l_e_st_eq_landau_n_rt_rp_negdirp r)). Time Defined. (* constant 3383 *) Definition l_e_st_eq_landau_n_rt_rp_remark4a : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), l_e_st_eq_landau_n_rt_rp_irratd (l_e_st_eq_landau_n_rt_rp_pdofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_ratd (l_e_st_eq_landau_n_rt_rp_pdofrp r)) (l_e_st_eq_landau_n_rt_rp_ratrp r) i (fun (t:l_e_st_eq_landau_n_rt_rp_ratd (l_e_st_eq_landau_n_rt_rp_pdofrp r)) => l_e_st_eq_landau_n_rt_rp_iv2d_t24 r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_pdofrp r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t18 r))) (l_e_st_eq_landau_n_rt_rp_iv2d_t20 r (l_e_st_eq_landau_n_rt_rp_iv2d_t18 r)) (t (l_e_st_eq_landau_n_rt_rp_iv2d_t18 r))))). Time Defined. (* constant 3384 *) Definition l_e_st_eq_landau_n_rt_rp_remark4b : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => l_e_st_eq_landau_n_rt_rp_iv2d_t17 r)). Time Defined. (* constant 3385 *) Definition l_e_st_eq_landau_n_rt_rp_remark5a : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), l_e_st_eq_landau_n_rt_rp_irratd (l_e_st_eq_landau_n_rt_rp_ndofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_ratd (l_e_st_eq_landau_n_rt_rp_ndofrp r)) (l_e_st_eq_landau_n_rt_rp_ratrp r) i (fun (t:l_e_st_eq_landau_n_rt_rp_ratd (l_e_st_eq_landau_n_rt_rp_ndofrp r)) => l_e_st_eq_landau_n_rt_rp_iv2d_t24 r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_ndofrp r)) (l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_ndofrp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t19 r))) (l_e_st_eq_landau_n_rt_rp_iv2d_t22 r (l_e_st_eq_landau_n_rt_rp_iv2d_t19 r)) (t (l_e_st_eq_landau_n_rt_rp_iv2d_t19 r))))). Time Defined. (* constant 3386 *) Definition l_e_st_eq_landau_n_rt_rp_remark5b : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_irratrp r), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_ndofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_irratrp r) => l_e_st_eq_landau_n_rt_rp_negdirp r)). Time Defined. (* constant 3387 *) Definition l_e_st_eq_landau_n_rt_rp_natd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_posd a) (forall (t:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd a t))). Time Defined. (* constant 3388 *) Definition l_e_st_eq_landau_n_rt_rp_natposd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), l_e_st_eq_landau_n_rt_rp_posd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => l_ande1 (l_e_st_eq_landau_n_rt_rp_posd a) (forall (t:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd a t)) n)). Time Defined. (* constant 3389 *) Definition l_e_st_eq_landau_n_rt_rp_natderp : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd a (l_e_st_eq_landau_n_rt_rp_natposd a n)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => l_r_ande2 (l_e_st_eq_landau_n_rt_rp_posd a) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd a t)) n)). Time Defined. (* constant 3390 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t25 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), l_e_st_eq_landau_n_rt_rp_posd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => l_e_st_eq_landau_n_rt_rp_eqposd a b e (l_e_st_eq_landau_n_rt_rp_natposd a n))))). Time Defined. (* constant 3391 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t26 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), (forall (p:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd a (l_e_st_eq_landau_n_rt_rp_natposd a n)) (l_e_st_eq_landau_n_rt_rp_rpofpd b p)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_eqpderp a (l_e_st_eq_landau_n_rt_rp_natposd a n) b p e))))). Time Defined. (* constant 3392 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t27 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), (forall (p:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd b p)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) (l_e_st_eq_landau_n_rt_rp_rpofpd a (l_e_st_eq_landau_n_rt_rp_natposd a n)) (l_e_st_eq_landau_n_rt_rp_rpofpd b p) (l_e_st_eq_landau_n_rt_rp_natderp a n) (l_e_st_eq_landau_n_rt_rp_iv2d_t26 a b e n p)))))). Time Defined. (* constant 3393 *) Definition l_e_st_eq_landau_n_rt_rp_eqnatd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), l_e_st_eq_landau_n_rt_rp_natd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => l_andi (l_e_st_eq_landau_n_rt_rp_posd b) (forall (t:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd b t)) (l_e_st_eq_landau_n_rt_rp_iv2d_t25 a b e n) (fun (t:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_iv2d_t27 a b e n t))))). Time Defined. (* constant 3394 *) Definition l_e_st_eq_landau_n_rt_rp_pdofnt : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_dif). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_rpofnt x)). Time Defined. (* constant 3395 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t28 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_rpofnt x)). Time Defined. (* constant 3396 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t29 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) (l_e_st_eq_landau_n_rt_rp_iv2d_t28 x)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)) => l_e_st_eq_landau_n_rt_rp_isrppd1 (l_e_st_eq_landau_n_rt_rp_rpofnt x))). Time Defined. (* constant 3397 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t30 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) (l_e_st_eq_landau_n_rt_rp_iv2d_t28 x)) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) p))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)) => l_e_st_eq_landau_n_rt_rp_eqpderp (l_e_st_eq_landau_n_rt_rp_pdofnt x) (l_e_st_eq_landau_n_rt_rp_iv2d_t28 x) (l_e_st_eq_landau_n_rt_rp_pdofnt x) p (l_e_st_eq_landau_n_rt_rp_refeq (l_e_st_eq_landau_n_rt_rp_pdofnt x)))). Time Defined. (* constant 3398 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t31 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) p))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) (l_e_st_eq_landau_n_rt_rp_iv2d_t28 x)) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) p) (l_e_st_eq_landau_n_rt_rp_iv2d_t29 x p) (l_e_st_eq_landau_n_rt_rp_iv2d_t30 x p))). Time Defined. (* constant 3399 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t32 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) p))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) p) (l_e_st_eq_landau_n_rt_rp_natrpi x) (l_e_st_eq_landau_n_rt_rp_iv2d_t31 x p))). Time Defined. (* constant 3400 *) Definition l_e_st_eq_landau_n_rt_rp_natdi : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_pdofnt x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_andi (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)) (forall (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofnt x) t)) (l_e_st_eq_landau_n_rt_rp_iv2d_t28 x) (fun (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofnt x)) => l_e_st_eq_landau_n_rt_rp_iv2d_t32 x t)). Time Defined. (* constant 3401 *) Definition l_e_st_eq_landau_n_rt_rp_intd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_or (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a))). Time Defined. (* constant 3402 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t33 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_zero b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_eqzero a b e z))))). Time Defined. (* constant 3403 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t34 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (n:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) => l_e_st_eq_landau_n_rt_rp_eqnatd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_eqabsd a b e) n))))). Time Defined. (* constant 3404 *) Definition l_e_st_eq_landau_n_rt_rp_eqintd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), l_e_st_eq_landau_n_rt_rp_intd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd b)) i (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_iv2d_t33 a b e i t) (fun (t:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) => l_e_st_eq_landau_n_rt_rp_iv2d_t34 a b e i t))))). Time Defined. (* constant 3405 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t34a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a (l_e_st_eq_landau_n_rt_rp_natposd a n))))). Time Defined. (* constant 3406 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t35 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => l_e_st_eq_landau_n_rt_rp_eqnatd a (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_iv2d_t34a a n) n)). Time Defined. (* constant 3407 *) Definition l_e_st_eq_landau_n_rt_rp_natintd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a), l_e_st_eq_landau_n_rt_rp_intd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a) => l_ori2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_iv2d_t35 a n))). Time Defined. (* constant 3408 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t36 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => l_ore2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) i (l_e_st_eq_landau_n_rt_rp_pnot0d a p)))). Time Defined. (* constant 3409 *) Definition l_e_st_eq_landau_n_rt_rp_posintnatd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), l_e_st_eq_landau_n_rt_rp_natd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => l_e_st_eq_landau_n_rt_rp_eqnatd (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a p)) (l_e_st_eq_landau_n_rt_rp_iv2d_t36 a p i)))). Time Defined. (* constant 3410 *) Definition l_e_st_eq_landau_n_rt_rp_intdi0 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_intd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_ori1 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) z)). Time Defined. (* constant 3411 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t37 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r) => l_e_st_eq_landau_n_rt_rp_posdirp r)). Time Defined. (* constant 3412 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t38 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r)), l_e_st_eq_landau_n_rt_rp_is r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) p)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r)) => l_e_tris l_e_st_eq_landau_n_rt_cut r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t37 r n)) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) p) (l_e_st_eq_landau_n_rt_rp_isrppd1 r) (l_e_st_eq_landau_n_rt_rp_eqpderp (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t37 r n) (l_e_st_eq_landau_n_rt_rp_pdofrp r) p (l_e_st_eq_landau_n_rt_rp_refeq (l_e_st_eq_landau_n_rt_rp_pdofrp r)))))). Time Defined. (* constant 3413 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t39 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r)), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) p)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r)) => l_e_isp l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) r (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) p) n (l_e_st_eq_landau_n_rt_rp_iv2d_t38 r n p)))). Time Defined. (* constant 3414 *) Definition l_e_st_eq_landau_n_rt_rp_remark6a : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_pdofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r) => l_andi (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r)) (forall (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r)), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp r) t)) (l_e_st_eq_landau_n_rt_rp_iv2d_t37 r n) (fun (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp r)) => l_e_st_eq_landau_n_rt_rp_iv2d_t39 r n t))). Time Defined. (* constant 3415 *) Definition l_e_st_eq_landau_n_rt_rp_remark6 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pdofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r) => l_e_st_eq_landau_n_rt_rp_natintd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_remark6a r n))). Time Defined. (* constant 3416 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t40 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_ndofrp r)) (l_e_st_eq_landau_n_rt_rp_pdofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r) => l_e_st_eq_landau_n_rt_rp_absdeql l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_isless2 (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp r) (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_compl l_e_st_eq_landau_n_rt_rp_1rp r) (l_e_st_eq_landau_n_rt_rp_satz133a l_e_st_eq_landau_n_rt_rp_1rp r)))). Time Defined. (* constant 3417 *) Definition l_e_st_eq_landau_n_rt_rp_iv2d_t41 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_ndofrp r)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r) => l_e_st_eq_landau_n_rt_rp_eqnatd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_ndofrp r)) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_ndofrp r)) (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_iv2d_t40 r n)) (l_e_st_eq_landau_n_rt_rp_remark6a r n))). Time Defined. (* constant 3418 *) Definition l_e_st_eq_landau_n_rt_rp_remark7 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_ndofrp r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r) => l_ori2 (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_ndofrp r)) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_ndofrp r))) (l_e_st_eq_landau_n_rt_rp_iv2d_t41 r n))). Time Defined. (* constant 3419 *) Definition l_e_st_eq_landau_n_rt_rp_2d174_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_ore2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) i n))). Time Defined. (* constant 3420 *) Definition l_e_st_eq_landau_n_rt_rp_2d174_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a)), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) t))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a)) (forall (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a)), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) t)) (l_e_st_eq_landau_n_rt_rp_2d174_t1 a i n)))). Time Defined. (* constant 3421 *) Definition l_e_st_eq_landau_n_rt_rp_2d174_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_ratrp (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_satzd166e a n))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_lemmaiii5 (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_satzd166e a n)) (l_e_st_eq_landau_n_rt_rp_2d174_t2 a i n (l_e_st_eq_landau_n_rt_rp_satzd166e a n))))). Time Defined. (* constant 3422 *) Definition l_e_st_eq_landau_n_rt_rp_satzd174 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), l_e_st_eq_landau_n_rt_rp_ratd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_2d174_t3 a i t))). Time Defined. (* constant 3423 *) Definition l_e_st_eq_landau_n_rt_rp_pd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_dif)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))). Time Defined. (* constant 3424 *) Definition l_e_st_eq_landau_n_rt_rp_pd12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b1) (l_e_st_eq_landau_n_rt_rp_pl a2 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_issmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_pl a1 b1) (l_e_st_eq_landau_n_rt_rp_pl a2 b2) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)) b1 (l_e_st_eq_landau_n_rt_rp_stmis a1 a2) (l_e_st_eq_landau_n_rt_rp_stmis b1 b2)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2)) b2 (l_e_st_eq_landau_n_rt_rp_stdis a1 a2) (l_e_st_eq_landau_n_rt_rp_stdis b1 b2)))))). Time Defined. (* constant 3425 *) Definition l_e_st_eq_landau_n_rt_rp_pd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r2))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_issmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r2) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) r1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_stmis r1 r2)) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2)) r2 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_stdis r1 r2))))). Time Defined. (* constant 3426 *) Definition l_e_st_eq_landau_n_rt_rp_pd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl r2 (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_issmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl r2 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) r1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_stmis r1 r2)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2)) r2 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_stdis r1 r2))))). Time Defined. (* constant 3427 *) Definition l_e_st_eq_landau_n_rt_rp_pdeq12a : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b1) (l_e_st_eq_landau_n_rt_rp_pl a2 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b1) (l_e_st_eq_landau_n_rt_rp_pl a2 b2)) (l_e_st_eq_landau_n_rt_rp_pd12 a1 a2 b1 b2))))). Time Defined. (* constant 3428 *) Definition l_e_st_eq_landau_n_rt_rp_pdeq12b : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b1) (l_e_st_eq_landau_n_rt_rp_pl a2 b2)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq2 (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b1) (l_e_st_eq_landau_n_rt_rp_pl a2 b2)) (l_e_st_eq_landau_n_rt_rp_pd12 a1 a2 b1 b2))))). Time Defined. (* constant 3429 *) Definition l_e_st_eq_landau_n_rt_rp_pdeq1a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r2))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pd1 a r1 r2)))). Time Defined. (* constant 3430 *) Definition l_e_st_eq_landau_n_rt_rp_pdeq1b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_df r1 r2))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq2 (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pd1 a r1 r2)))). Time Defined. (* constant 3431 *) Definition l_e_st_eq_landau_n_rt_rp_pdeq2a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl r2 (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pd2 a r1 r2)))). Time Defined. (* constant 3432 *) Definition l_e_st_eq_landau_n_rt_rp_pdeq2b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df r1 r2) a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq2 (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pd2 a r1 r2)))). Time Defined. (* constant 3433 *) Definition l_e_st_eq_landau_n_rt_rp_satzd175 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pd b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))). Time Defined. (* constant 3434 *) Definition l_e_st_eq_landau_n_rt_rp_compd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pd b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satzd175 a b)). Time Defined. (* constant 3435 *) Definition l_e_st_eq_landau_n_rt_rp_iv3d_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_4pl23 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2c a b c)) e) (l_e_st_eq_landau_n_rt_rp_4pl23 (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2c a b c)))))). Time Defined. (* constant 3436 *) Definition l_e_st_eq_landau_n_rt_rp_eqpd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_iv3d_t1 a b c e))))). Time Defined. (* constant 3437 *) Definition l_e_st_eq_landau_n_rt_rp_eqpd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_pd c b) (l_e_st_eq_landau_n_rt_rp_compd c a) (l_e_st_eq_landau_n_rt_rp_eqpd1 a b c e) (l_e_st_eq_landau_n_rt_rp_compd b c))))). Time Defined. (* constant 3438 *) Definition l_e_st_eq_landau_n_rt_rp_eqpd12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_eqpd1 a b c e) (l_e_st_eq_landau_n_rt_rp_eqpd2 c d b f))))))). Time Defined. (* constant 3439 *) Definition l_e_st_eq_landau_n_rt_rp_iv3d_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_tr4is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b)) z (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_asspl2 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_1b a b))))). Time Defined. (* constant 3440 *) Definition l_e_st_eq_landau_n_rt_rp_pd01 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a b) b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_eqi2 b (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_iv3d_t2 a b z)))). Time Defined. (* constant 3441 *) Definition l_e_st_eq_landau_n_rt_rp_pd02 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a b) a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pd b a) a (l_e_st_eq_landau_n_rt_rp_compd a b) (l_e_st_eq_landau_n_rt_rp_pd01 b a z)))). Time Defined. (* constant 3442 *) Definition l_e_st_eq_landau_n_rt_rp_ppd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_posdi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_satz137 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) p q))))). Time Defined. (* constant 3443 *) Definition l_e_st_eq_landau_n_rt_rp_npd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (o:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_negdi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_satz137a (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) n o))))). Time Defined. (* constant 3444 *) Definition l_e_st_eq_landau_n_rt_rp_m0d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_dif). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)). Time Defined. (* constant 3445 *) Definition l_e_st_eq_landau_n_rt_rp_m0deqa : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_df a2 a1))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2 a1 (l_e_st_eq_landau_n_rt_rp_stdis a1 a2) (l_e_st_eq_landau_n_rt_rp_stmis a1 a2))). Time Defined. (* constant 3446 *) Definition l_e_st_eq_landau_n_rt_rp_m0deqb : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df a2 a1) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_df a1 a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_df a2 a1) (l_e_st_eq_landau_n_rt_rp_m0deqa a1 a2))). Time Defined. (* constant 3447 *) Definition l_e_st_eq_landau_n_rt_rp_iv3d_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_symeq a b e) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))))). Time Defined. (* constant 3448 *) Definition l_e_st_eq_landau_n_rt_rp_eqm0d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_iv3d_t3 a b e)))). Time Defined. (* constant 3449 *) Definition l_e_st_eq_landau_n_rt_rp_satzd176a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_negdi (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) p))). Time Defined. (* constant 3450 *) Definition l_e_st_eq_landau_n_rt_rp_satzd176b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_zeroi (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) z))). Time Defined. (* constant 3451 *) Definition l_e_st_eq_landau_n_rt_rp_satzd176c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_posdi (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) n))). Time Defined. (* constant 3452 *) Definition l_e_st_eq_landau_n_rt_rp_satzd176d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_m0d a)), l_e_st_eq_landau_n_rt_rp_posd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_m0d a)) => l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_stmis (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_stdis (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) n))). Time Defined. (* constant 3453 *) Definition l_e_st_eq_landau_n_rt_rp_satzd176e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_m0d a)), l_e_st_eq_landau_n_rt_rp_zero a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_m0d a)) => l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_isstm (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) z (l_e_st_eq_landau_n_rt_rp_stdis (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a))))). Time Defined. (* constant 3454 *) Definition l_e_st_eq_landau_n_rt_rp_satzd176f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_m0d a)), l_e_st_eq_landau_n_rt_rp_negd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_m0d a)) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_stmis (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_stdis (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) p))). Time Defined. (* constant 3455 *) Definition l_e_st_eq_landau_n_rt_rp_m0d0 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_zeroeq (l_e_st_eq_landau_n_rt_rp_m0d a) a (l_e_st_eq_landau_n_rt_rp_satzd176b a z) z)). Time Defined. (* constant 3456 *) Definition l_e_st_eq_landau_n_rt_rp_3d177_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) a). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_tris l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) a (l_e_st_eq_landau_n_rt_rp_issmsd (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_stdis (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_stmis (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_dfis a)). Time Defined. (* constant 3457 *) Definition l_e_st_eq_landau_n_rt_rp_satzd177 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) a). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) a (l_e_st_eq_landau_n_rt_rp_3d177_t1 a)). Time Defined. (* constant 3458 *) Definition l_e_st_eq_landau_n_rt_rp_satzd177a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) a (l_e_st_eq_landau_n_rt_rp_satzd177 a)). Time Defined. (* constant 3459 *) Definition l_e_st_eq_landau_n_rt_rp_satzd177b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_m0d b)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_m0d b)) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b)) b (l_e_st_eq_landau_n_rt_rp_eqm0d a (l_e_st_eq_landau_n_rt_rp_m0d b) e) (l_e_st_eq_landau_n_rt_rp_satzd177 b)))). Time Defined. (* constant 3460 *) Definition l_e_st_eq_landau_n_rt_rp_satzd177c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_m0d b)), l_e_st_eq_landau_n_rt_rp_eq b (l_e_st_eq_landau_n_rt_rp_m0d a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_m0d b)) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_m0d a) b (l_e_st_eq_landau_n_rt_rp_satzd177b a b e)))). Time Defined. (* constant 3461 *) Definition l_e_st_eq_landau_n_rt_rp_satzd177d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) b), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) b) => l_e_st_eq_landau_n_rt_rp_satzd177c b a (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_m0d a) b e)))). Time Defined. (* constant 3462 *) Definition l_e_st_eq_landau_n_rt_rp_satzd177e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d b) a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) b) => l_e_st_eq_landau_n_rt_rp_symeq a (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_satzd177d a b e)))). Time Defined. (* constant 3463 *) Definition l_e_st_eq_landau_n_rt_rp_3d178_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) a (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absnd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_satzd176a a p)) (l_e_st_eq_landau_n_rt_rp_satzd177 a) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a p))))). Time Defined. (* constant 3464 *) Definition l_e_st_eq_landau_n_rt_rp_3d178_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_m0d a) a (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absnnd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_0notnd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_satzd176b a z))) (l_e_st_eq_landau_n_rt_rp_m0d0 a z) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_0notnd a z))))). Time Defined. (* constant 3465 *) Definition l_e_st_eq_landau_n_rt_rp_3d178_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absnnd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_pnotnd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_satzd176c a n))) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_absnd a n)))). Time Defined. (* constant 3466 *) Definition l_e_st_eq_landau_n_rt_rp_satzd178 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_absd a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_rappd a (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_absd a)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_3d178_t1 a t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_3d178_t2 a t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_3d178_t3 a t)). Time Defined. (* constant 3467 *) Definition l_e_st_eq_landau_n_rt_rp_satzd178a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_satzd178 a)). Time Defined. (* constant 3468 *) Definition l_e_st_eq_landau_n_rt_rp_3d179_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_pdeq1b a (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)). Time Defined. (* constant 3469 *) Definition l_e_st_eq_landau_n_rt_rp_3d179_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_zeroi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a))). Time Defined. (* constant 3470 *) Definition l_e_st_eq_landau_n_rt_rp_satzd179 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_3d179_t1 a) (l_e_st_eq_landau_n_rt_rp_3d179_t2 a)). Time Defined. (* constant 3471 *) Definition l_e_st_eq_landau_n_rt_rp_satzd179a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) a) (l_e_st_eq_landau_n_rt_rp_compd a (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_satzd179 a)). Time Defined. (* constant 3472 *) Definition l_e_st_eq_landau_n_rt_rp_satzd180 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_m0deqa (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pdeq12b (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b)))). Time Defined. (* constant 3473 *) Definition l_e_st_eq_landau_n_rt_rp_satzd180a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_satzd180 a b))). Time Defined. (* constant 3474 *) Definition l_e_st_eq_landau_n_rt_rp_md : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_dif)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_m0d b))). Time Defined. (* constant 3475 *) Definition l_e_st_eq_landau_n_rt_rp_mdeq12a : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl a2 b1)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b2 b1)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl a2 b1)) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df b2 b1) (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_m0deqa b1 b2)) (l_e_st_eq_landau_n_rt_rp_pdeq12a a1 a2 b2 b1))))). Time Defined. (* constant 3476 *) Definition l_e_st_eq_landau_n_rt_rp_mdeq12b : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl a2 b1)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl a1 b2) (l_e_st_eq_landau_n_rt_rp_pl a2 b1)) (l_e_st_eq_landau_n_rt_rp_mdeq12a a1 a2 b1 b2))))). Time Defined. (* constant 3477 *) Definition l_e_st_eq_landau_n_rt_rp_mdeq1a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r1))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_md a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_df r2 r1)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r1)) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df r2 r1) a (l_e_st_eq_landau_n_rt_rp_m0deqa r1 r2)) (l_e_st_eq_landau_n_rt_rp_pdeq1a a r2 r1)))). Time Defined. (* constant 3478 *) Definition l_e_st_eq_landau_n_rt_rp_mdeq1b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r1)) (l_e_st_eq_landau_n_rt_rp_md a (l_e_st_eq_landau_n_rt_rp_df r1 r2))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_md a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) r1)) (l_e_st_eq_landau_n_rt_rp_mdeq1a a r1 r2)))). Time Defined. (* constant 3479 *) Definition l_e_st_eq_landau_n_rt_rp_mdeq2a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl r2 (l_e_st_eq_landau_n_rt_rp_1a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pdeq12a r1 r2 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)))). Time Defined. (* constant 3480 *) Definition l_e_st_eq_landau_n_rt_rp_mdeq2b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl r2 (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_df r1 r2) a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pdeq12b r1 r2 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)))). Time Defined. (* constant 3481 *) Definition l_e_st_eq_landau_n_rt_rp_eqmd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md a c) (l_e_st_eq_landau_n_rt_rp_md b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqpd1 a b (l_e_st_eq_landau_n_rt_rp_m0d c) e)))). Time Defined. (* constant 3482 *) Definition l_e_st_eq_landau_n_rt_rp_eqmd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md c a) (l_e_st_eq_landau_n_rt_rp_md c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b) c (l_e_st_eq_landau_n_rt_rp_eqm0d a b e))))). Time Defined. (* constant 3483 *) Definition l_e_st_eq_landau_n_rt_rp_eqmd12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md a c) (l_e_st_eq_landau_n_rt_rp_md b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_md a c) (l_e_st_eq_landau_n_rt_rp_md b c) (l_e_st_eq_landau_n_rt_rp_md b d) (l_e_st_eq_landau_n_rt_rp_eqmd1 a b c e) (l_e_st_eq_landau_n_rt_rp_eqmd2 c d b f))))))). Time Defined. (* constant 3484 *) Definition l_e_st_eq_landau_n_rt_rp_satzd181 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_md b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b))) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) b) (l_e_st_eq_landau_n_rt_rp_md b a) (l_e_st_eq_landau_n_rt_rp_satzd180 a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b)) b (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_satzd177 b)) (l_e_st_eq_landau_n_rt_rp_compd (l_e_st_eq_landau_n_rt_rp_m0d a) b))). Time Defined. (* constant 3485 *) Definition l_e_st_eq_landau_n_rt_rp_satzd181a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_md b a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_md b a)) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_satzd181 b a))). Time Defined. (* constant 3486 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pdeq1a a (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_eqsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))))). Time Defined. (* constant 3487 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_md a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_3d182_t1 a b))). Time Defined. (* constant 3488 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_stmis (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3489 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_stdis (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3490 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a b)), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a b)) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_3d182_t1 a b) p))). Time Defined. (* constant 3491 *) Definition l_e_st_eq_landau_n_rt_rp_satzd182a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a b)), l_e_st_eq_landau_n_rt_rp_mored a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a b)) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_3d182_t3 a b) (l_e_st_eq_landau_n_rt_rp_3d182_t4 a b) (l_e_st_eq_landau_n_rt_rp_3d182_t5 a b p)))). Time Defined. (* constant 3492 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b)), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b)) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_3d182_t1 a b) z))). Time Defined. (* constant 3493 *) Definition l_e_st_eq_landau_n_rt_rp_satzd182b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b)), l_e_st_eq_landau_n_rt_rp_eq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b)) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_isstm (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_3d182_t6 a b z) (l_e_st_eq_landau_n_rt_rp_3d182_t4 a b)))). Time Defined. (* constant 3494 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md a b)), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md a b)) => l_e_st_eq_landau_n_rt_rp_eqnegd (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_3d182_t1 a b) n))). Time Defined. (* constant 3495 *) Definition l_e_st_eq_landau_n_rt_rp_satzd182c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md a b)), l_e_st_eq_landau_n_rt_rp_lessd a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md a b)) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_3d182_t3 a b) (l_e_st_eq_landau_n_rt_rp_3d182_t4 a b) (l_e_st_eq_landau_n_rt_rp_3d182_t7 a b n)))). Time Defined. (* constant 3496 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t8 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_posdi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) m))). Time Defined. (* constant 3497 *) Definition l_e_st_eq_landau_n_rt_rp_satzd182d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_3d182_t2 a b) (l_e_st_eq_landau_n_rt_rp_3d182_t8 a b m)))). Time Defined. (* constant 3498 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t9 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_zeroi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) e))). Time Defined. (* constant 3499 *) Definition l_e_st_eq_landau_n_rt_rp_satzd182e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_3d182_t2 a b) (l_e_st_eq_landau_n_rt_rp_3d182_t9 a b e)))). Time Defined. (* constant 3500 *) Definition l_e_st_eq_landau_n_rt_rp_3d182_t10 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_negdi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) l))). Time Defined. (* constant 3501 *) Definition l_e_st_eq_landau_n_rt_rp_satzd182f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_eqnegd (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_3d182_t2 a b) (l_e_st_eq_landau_n_rt_rp_3d182_t10 a b l)))). Time Defined. (* constant 3502 *) Definition l_e_st_eq_landau_n_rt_rp_3d183_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d a))) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_12issmsd (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)))). Time Defined. (* constant 3503 *) Definition l_e_st_eq_landau_n_rt_rp_3d183_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_3d183_t1 b a)). Time Defined. (* constant 3504 *) Definition l_e_st_eq_landau_n_rt_rp_satzd183a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_isless12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d a))) (l_e_st_eq_landau_n_rt_rp_3d183_t2 a b) (l_e_st_eq_landau_n_rt_rp_3d183_t1 a b) (l_e_st_eq_landau_n_rt_rp_lemmad5 a b m)))). Time Defined. (* constant 3505 *) Definition l_e_st_eq_landau_n_rt_rp_staz183b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqm0d a b e))). Time Defined. (* constant 3506 *) Definition l_e_st_eq_landau_n_rt_rp_satzd183c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_m0d a))) (l_e_st_eq_landau_n_rt_rp_3d183_t2 a b) (l_e_st_eq_landau_n_rt_rp_3d183_t1 a b) (l_e_st_eq_landau_n_rt_rp_lemmad6 a b l)))). Time Defined. (* constant 3507 *) Definition l_e_st_eq_landau_n_rt_rp_satzd183d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)), l_e_st_eq_landau_n_rt_rp_mored a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) => l_e_st_eq_landau_n_rt_rp_eqmored12 (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) a (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b)) b (l_e_st_eq_landau_n_rt_rp_satzd177 a) (l_e_st_eq_landau_n_rt_rp_satzd177 b) (l_e_st_eq_landau_n_rt_rp_satzd183c (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b) l)))). Time Defined. (* constant 3508 *) Definition l_e_st_eq_landau_n_rt_rp_satzd183e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)), l_e_st_eq_landau_n_rt_rp_eq a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) => l_e_st_eq_landau_n_rt_rp_tr3eq a (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b)) b (l_e_st_eq_landau_n_rt_rp_satzd177a a) (l_e_st_eq_landau_n_rt_rp_eqm0d (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b) e) (l_e_st_eq_landau_n_rt_rp_satzd177 b)))). Time Defined. (* constant 3509 *) Definition l_e_st_eq_landau_n_rt_rp_satzd183f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)), l_e_st_eq_landau_n_rt_rp_lessd a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) => l_e_st_eq_landau_n_rt_rp_eqlessd12 (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a)) a (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b)) b (l_e_st_eq_landau_n_rt_rp_satzd177 a) (l_e_st_eq_landau_n_rt_rp_satzd177 b) (l_e_st_eq_landau_n_rt_rp_satzd183a (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b) m)))). Time Defined. (* constant 3510 *) Definition l_e_st_eq_landau_n_rt_rp_3d184_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_2a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq a (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_lemmad3 a (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_asspl2 (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_3pl12 (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_mdeq12b (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp)). Time Defined. (* constant 3511 *) Definition l_e_st_eq_landau_n_rt_rp_3d184_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_and3 (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_2a a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_and3i (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_3d184_t1 a)). Time Defined. (* constant 3512 *) Definition l_e_st_eq_landau_n_rt_rp_3d184_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_posd x) (l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a)) x)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_posd x) (l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a)) x))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_3d184_t2 a)). Time Defined. (* constant 3513 *) Definition l_e_st_eq_landau_n_rt_rp_satzd184 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_posd x) (l_e_st_eq_landau_n_rt_rp_posd y) (l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_md x y))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_posd x) (l_e_st_eq_landau_n_rt_rp_posd y) (l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_md x y)))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_3d184_t3 a)). Time Defined. (* constant 3514 *) Definition l_e_st_eq_landau_n_rt_rp_asspd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_pdeq2a c (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pdeq1b a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)))))). Time Defined. (* constant 3515 *) Definition l_e_st_eq_landau_n_rt_rp_asspd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_asspd1 a b c)))). Time Defined. (* constant 3516 *) Definition l_e_st_eq_landau_n_rt_rp_3pd23 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a c) b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd c b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a c) b) (l_e_st_eq_landau_n_rt_rp_asspd1 a b c) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_pd c b) a (l_e_st_eq_landau_n_rt_rp_compd b c)) (l_e_st_eq_landau_n_rt_rp_asspd2 a c b)))). Time Defined. (* constant 3517 *) Definition l_e_st_eq_landau_n_rt_rp_4pd23 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pd c d)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pd c d)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c) d) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a c) b) d) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)) (l_e_st_eq_landau_n_rt_rp_asspd2 (l_e_st_eq_landau_n_rt_rp_pd a b) c d) (l_e_st_eq_landau_n_rt_rp_eqpd1 (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a c) b) d (l_e_st_eq_landau_n_rt_rp_3pd23 a b c)) (l_e_st_eq_landau_n_rt_rp_asspd1 (l_e_st_eq_landau_n_rt_rp_pd a c) b d))))). Time Defined. (* constant 3518 *) Definition l_e_st_eq_landau_n_rt_rp_pdmd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md a b) b) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md a b) b) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d b) b)) a (l_e_st_eq_landau_n_rt_rp_asspd1 a (l_e_st_eq_landau_n_rt_rp_m0d b) b) (l_e_st_eq_landau_n_rt_rp_pd02 a (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d b) b) (l_e_st_eq_landau_n_rt_rp_satzd179a b)))). Time Defined. (* constant 3519 *) Definition l_e_st_eq_landau_n_rt_rp_mdpd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a b) b) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a b) b) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd b (l_e_st_eq_landau_n_rt_rp_m0d b))) a (l_e_st_eq_landau_n_rt_rp_asspd1 a b (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_pd02 a (l_e_st_eq_landau_n_rt_rp_pd b (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_satzd179 b)))). Time Defined. (* constant 3520 *) Definition l_e_st_eq_landau_n_rt_rp_satzd185 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_md c d)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_md c d)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_m0d d))) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)) (l_e_st_eq_landau_n_rt_rp_4pd23 a (l_e_st_eq_landau_n_rt_rp_m0d b) c (l_e_st_eq_landau_n_rt_rp_m0d d)) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_m0d d)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd b d)) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_satzd180a b d)))))). Time Defined. (* constant 3521 *) Definition l_e_st_eq_landau_n_rt_rp_satzd186 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_pd a (l_e_st_eq_landau_n_rt_rp_pd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_asspd1 a b c))). Time Defined. (* constant 3522 *) Definition l_e_st_eq_landau_n_rt_rp_satzd187a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd b (l_e_st_eq_landau_n_rt_rp_md a b)) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd b (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md a b) b) a (l_e_st_eq_landau_n_rt_rp_compd b (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_pdmd a b))). Time Defined. (* constant 3523 *) Definition l_e_st_eq_landau_n_rt_rp_satzd187c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd b x) a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md a b) x)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd b x) a) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd x b) b) x (l_e_st_eq_landau_n_rt_rp_eqmd1 a (l_e_st_eq_landau_n_rt_rp_pd x b) b (l_e_st_eq_landau_n_rt_rp_treq1 a (l_e_st_eq_landau_n_rt_rp_pd x b) (l_e_st_eq_landau_n_rt_rp_pd b x) e (l_e_st_eq_landau_n_rt_rp_compd b x))) (l_e_st_eq_landau_n_rt_rp_mdpd x b))))). Time Defined. (* constant 3524 *) Definition l_e_st_eq_landau_n_rt_rp_satzd187d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd b x) a), l_e_st_eq_landau_n_rt_rp_eq x (l_e_st_eq_landau_n_rt_rp_md a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd b x) a) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_md a b) x (l_e_st_eq_landau_n_rt_rp_satzd187c a b x e))))). Time Defined. (* constant 3525 *) Definition l_e_st_eq_landau_n_rt_rp_satzd187e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd x b) a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md a b) x)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd x b) a) => l_e_st_eq_landau_n_rt_rp_satzd187c a b x (l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd b x) (l_e_st_eq_landau_n_rt_rp_pd x b) a (l_e_st_eq_landau_n_rt_rp_compd b x) e))))). Time Defined. (* constant 3526 *) Definition l_e_st_eq_landau_n_rt_rp_satzd187f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd x b) a), l_e_st_eq_landau_n_rt_rp_eq x (l_e_st_eq_landau_n_rt_rp_md a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd x b) a) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_md a b) x (l_e_st_eq_landau_n_rt_rp_satzd187e a b x e))))). Time Defined. (* constant 3527 *) Definition l_e_st_eq_landau_n_rt_rp_3d188_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_md a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_m0d c))) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_md c c)) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_m0d c)) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_satzd180 b c)) (l_e_st_eq_landau_n_rt_rp_4pd23 a c (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_m0d c)) (l_e_st_eq_landau_n_rt_rp_pd02 (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_md c c) (l_e_st_eq_landau_n_rt_rp_satzd179 c))))). Time Defined. (* constant 3528 *) Definition l_e_st_eq_landau_n_rt_rp_3d188_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_3d188_t1 a b c)))). Time Defined. (* constant 3529 *) Definition l_e_st_eq_landau_n_rt_rp_3d188_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_3d188_t1 a b c) (l_e_st_eq_landau_n_rt_rp_satzd182d (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c) m))))). Time Defined. (* constant 3530 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)), l_e_st_eq_landau_n_rt_rp_mored a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) => l_e_st_eq_landau_n_rt_rp_satzd182a a b (l_e_st_eq_landau_n_rt_rp_3d188_t3 a b c m))))). Time Defined. (* constant 3531 *) Definition l_e_st_eq_landau_n_rt_rp_3d188_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_3d188_t1 a b c) (l_e_st_eq_landau_n_rt_rp_satzd182e (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c) e))))). Time Defined. (* constant 3532 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)), l_e_st_eq_landau_n_rt_rp_eq a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) => l_e_st_eq_landau_n_rt_rp_satzd182b a b (l_e_st_eq_landau_n_rt_rp_3d188_t4 a b c e))))). Time Defined. (* constant 3533 *) Definition l_e_st_eq_landau_n_rt_rp_3d188_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) => l_e_st_eq_landau_n_rt_rp_eqnegd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_3d188_t1 a b c) (l_e_st_eq_landau_n_rt_rp_satzd182f (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c) l))))). Time Defined. (* constant 3534 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)), l_e_st_eq_landau_n_rt_rp_lessd a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) => l_e_st_eq_landau_n_rt_rp_satzd182c a b (l_e_st_eq_landau_n_rt_rp_3d188_t5 a b c l))))). Time Defined. (* constant 3535 *) Definition l_e_st_eq_landau_n_rt_rp_3d188_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_3d188_t2 a b c) (l_e_st_eq_landau_n_rt_rp_satzd182d a b m))))). Time Defined. (* constant 3536 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_satzd182a (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_3d188_t6 a b c m))))). Time Defined. (* constant 3537 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqpd1 a b c e)))). Time Defined. (* constant 3538 *) Definition l_e_st_eq_landau_n_rt_rp_3d188_t7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_eqnegd (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_3d188_t2 a b c) (l_e_st_eq_landau_n_rt_rp_satzd182f a b l))))). Time Defined. (* constant 3539 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_satzd182c (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_3d188_t7 a b c l))))). Time Defined. (* constant 3540 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188g : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b)), l_e_st_eq_landau_n_rt_rp_mored a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b)) => l_e_st_eq_landau_n_rt_rp_satzd188a a b c (l_e_st_eq_landau_n_rt_rp_eqmored12 (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd c b) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_compd c a) (l_e_st_eq_landau_n_rt_rp_compd c b) m))))). Time Defined. (* constant 3541 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188h : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b)), l_e_st_eq_landau_n_rt_rp_eq a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b)) => l_e_st_eq_landau_n_rt_rp_satzd188b a b c (l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_compd a c) e (l_e_st_eq_landau_n_rt_rp_compd c b)))))). Time Defined. (* constant 3542 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188j : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b)), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b)), l_e_st_eq_landau_n_rt_rp_lessd a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b)) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b)) => l_e_st_eq_landau_n_rt_rp_satzd188c a b c (l_e_st_eq_landau_n_rt_rp_eqlessd12 (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd c b) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_compd c a) (l_e_st_eq_landau_n_rt_rp_compd c b) l)))))). Time Defined. (* constant 3543 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188k : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_eqmored12 (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_pd c b) (l_e_st_eq_landau_n_rt_rp_compd a c) (l_e_st_eq_landau_n_rt_rp_compd b c) (l_e_st_eq_landau_n_rt_rp_satzd188d a b c m))))). Time Defined. (* constant 3544 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188l : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqpd2 a b c e)))). Time Defined. (* constant 3545 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188m : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => l_e_st_eq_landau_n_rt_rp_eqlessd12 (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_pd c b) (l_e_st_eq_landau_n_rt_rp_compd a c) (l_e_st_eq_landau_n_rt_rp_compd b c) (l_e_st_eq_landau_n_rt_rp_satzd188f a b c l))))). Time Defined. (* constant 3546 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188n : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (m:l_e_st_eq_landau_n_rt_rp_mored c d), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored c d) => l_e_st_eq_landau_n_rt_rp_eqmored2 (l_e_st_eq_landau_n_rt_rp_pd a d) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_eqpd1 a b d e) (l_e_st_eq_landau_n_rt_rp_satzd188k c d a m))))))). Time Defined. (* constant 3547 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188o : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (m:l_e_st_eq_landau_n_rt_rp_mored c d), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd d b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored c d) => l_e_st_eq_landau_n_rt_rp_eqmored12 (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_pd d b) (l_e_st_eq_landau_n_rt_rp_compd a c) (l_e_st_eq_landau_n_rt_rp_compd b d) (l_e_st_eq_landau_n_rt_rp_satzd188n a b c d e m))))))). Time Defined. (* constant 3548 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188p : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd c d), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd c d) => l_e_st_eq_landau_n_rt_rp_eqlessd2 (l_e_st_eq_landau_n_rt_rp_pd a d) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_eqpd1 a b d e) (l_e_st_eq_landau_n_rt_rp_satzd188m c d a l))))))). Time Defined. (* constant 3549 *) Definition l_e_st_eq_landau_n_rt_rp_satzd188q : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd c d), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd d b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd c d) => l_e_st_eq_landau_n_rt_rp_eqlessd12 (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_pd d b) (l_e_st_eq_landau_n_rt_rp_compd a c) (l_e_st_eq_landau_n_rt_rp_compd b d) (l_e_st_eq_landau_n_rt_rp_satzd188p a b c d e l))))))). Time Defined. (* constant 3550 *) Definition l_e_st_eq_landau_n_rt_rp_satzd189 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (n:l_e_st_eq_landau_n_rt_rp_mored c d), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_mored c d) => l_e_st_eq_landau_n_rt_rp_trmored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b c) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_satzd188d a b c m) (l_e_st_eq_landau_n_rt_rp_satzd188k c d b n))))))). Time Defined. (* constant 3551 *) Definition l_e_st_eq_landau_n_rt_rp_satzd189a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lessd c d), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessd c d) => l_e_st_eq_landau_n_rt_rp_lemmad5 (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_satzd189 b a d c (l_e_st_eq_landau_n_rt_rp_lemmad6 a b l) (l_e_st_eq_landau_n_rt_rp_lemmad6 c d k)))))))). Time Defined. (* constant 3552 *) Definition l_e_st_eq_landau_n_rt_rp_satzd190a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_mored c d), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_mored c d) => l_orapp (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)) m (fun (t:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_satzd189 a b c d t n) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_satzd188n a b c d t n))))))). Time Defined. (* constant 3553 *) Definition l_e_st_eq_landau_n_rt_rp_satzd190b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (n:l_e_st_eq_landau_n_rt_rp_moreq c d), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreq c d) => l_e_st_eq_landau_n_rt_rp_eqmored12 (l_e_st_eq_landau_n_rt_rp_pd c a) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd d b) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_compd c a) (l_e_st_eq_landau_n_rt_rp_compd d b) (l_e_st_eq_landau_n_rt_rp_satzd190a c d a b n m))))))). Time Defined. (* constant 3554 *) Definition l_e_st_eq_landau_n_rt_rp_satzd190c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lessd c d), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lessd c d) => l_e_st_eq_landau_n_rt_rp_lemmad5 (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_satzd190a b a d c (l_e_st_eq_landau_n_rt_rp_satzd168b a b l) (l_e_st_eq_landau_n_rt_rp_lemmad6 c d k)))))))). Time Defined. (* constant 3555 *) Definition l_e_st_eq_landau_n_rt_rp_satzd190d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lesseq c d), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lesseq c d) => l_e_st_eq_landau_n_rt_rp_lemmad5 (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_satzd190b b a d c (l_e_st_eq_landau_n_rt_rp_lemmad6 a b l) (l_e_st_eq_landau_n_rt_rp_satzd168b c d k)))))))). Time Defined. (* constant 3556 *) Definition l_e_st_eq_landau_n_rt_rp_3d191_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_moreq c d), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreq c d) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => l_e_st_eq_landau_n_rt_rp_moreqi2 (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_eqpd12 a b c d e f))))))))). Time Defined. (* constant 3557 *) Definition l_e_st_eq_landau_n_rt_rp_3d191_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_moreq c d), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (o:l_e_st_eq_landau_n_rt_rp_mored c d), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreq c d) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (o:l_e_st_eq_landau_n_rt_rp_mored c d) => l_e_st_eq_landau_n_rt_rp_moreqi1 (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_satzd190a a b c d m o))))))))). Time Defined. (* constant 3558 *) Definition l_e_st_eq_landau_n_rt_rp_3d191_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_moreq c d), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreq c d) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_orapp (l_e_st_eq_landau_n_rt_rp_mored c d) (l_e_st_eq_landau_n_rt_rp_eq c d) (l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)) n (fun (t:l_e_st_eq_landau_n_rt_rp_mored c d) => l_e_st_eq_landau_n_rt_rp_3d191_t2 a b c d m n e t) (fun (t:l_e_st_eq_landau_n_rt_rp_eq c d) => l_e_st_eq_landau_n_rt_rp_3d191_t1 a b c d m n e t)))))))). Time Defined. (* constant 3559 *) Definition l_e_st_eq_landau_n_rt_rp_3d191_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_moreq c d), (forall (o:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreq c d) => (fun (o:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_moreqi1 (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_satzd190b a b c d o n)))))))). Time Defined. (* constant 3560 *) Definition l_e_st_eq_landau_n_rt_rp_satzd191 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a b), (forall (n:l_e_st_eq_landau_n_rt_rp_moreq c d), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_moreq c d) => l_orapp (l_e_st_eq_landau_n_rt_rp_mored a b) (l_e_st_eq_landau_n_rt_rp_eq a b) (l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d)) m (fun (t:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_3d191_t4 a b c d m n t) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_3d191_t3 a b c d m n t))))))). Time Defined. (* constant 3561 *) Definition l_e_st_eq_landau_n_rt_rp_satzd191a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a b), (forall (k:l_e_st_eq_landau_n_rt_rp_lesseq c d), l_e_st_eq_landau_n_rt_rp_lesseq (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a b) => (fun (k:l_e_st_eq_landau_n_rt_rp_lesseq c d) => l_e_st_eq_landau_n_rt_rp_satzd168a (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_satzd191 b a d c (l_e_st_eq_landau_n_rt_rp_satzd168b a b l) (l_e_st_eq_landau_n_rt_rp_satzd168b c d k)))))))). Time Defined. (* constant 3562 *) Definition l_e_st_eq_landau_n_rt_rp_td : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_dif)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))))). Time Defined. (* constant 3563 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t1 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) r) (l_e_st_eq_landau_n_rt_rp_ts a1 r)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ists1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a1 r (l_e_st_eq_landau_n_rt_rp_stmis a1 a2)))). Time Defined. (* constant 3564 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t2 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_ts r a1)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ists2 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a1 r (l_e_st_eq_landau_n_rt_rp_stmis a1 a2)))). Time Defined. (* constant 3565 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t3 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) r) (l_e_st_eq_landau_n_rt_rp_ts a2 r)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ists1 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2 r (l_e_st_eq_landau_n_rt_rp_stdis a1 a2)))). Time Defined. (* constant 3566 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t4 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_ts r a2)))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ists2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) a2 r (l_e_st_eq_landau_n_rt_rp_stdis a1 a2)))). Time Defined. (* constant 3567 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t5 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 r) (l_e_st_eq_landau_n_rt_rp_ts a2 s)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) r) (l_e_st_eq_landau_n_rt_rp_ts a1 r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) s) (l_e_st_eq_landau_n_rt_rp_ts a2 s) (l_e_st_eq_landau_n_rt_rp_iv4d_t1 a1 a2 r) (l_e_st_eq_landau_n_rt_rp_iv4d_t3 a1 a2 s))))). Time Defined. (* constant 3568 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t6 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_ts s (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r a1) (l_e_st_eq_landau_n_rt_rp_ts s a2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_ts r a1) (l_e_st_eq_landau_n_rt_rp_ts s (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_ts s a2) (l_e_st_eq_landau_n_rt_rp_iv4d_t2 a1 a2 r) (l_e_st_eq_landau_n_rt_rp_iv4d_t4 a1 a2 s))))). Time Defined. (* constant 3569 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t7 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a2 r) (l_e_st_eq_landau_n_rt_rp_ts a1 s)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) r) (l_e_st_eq_landau_n_rt_rp_ts a2 r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) s) (l_e_st_eq_landau_n_rt_rp_ts a1 s) (l_e_st_eq_landau_n_rt_rp_iv4d_t3 a1 a2 r) (l_e_st_eq_landau_n_rt_rp_iv4d_t1 a1 a2 s))))). Time Defined. (* constant 3570 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t8 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_ts s (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r a2) (l_e_st_eq_landau_n_rt_rp_ts s a1)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_ts r a2) (l_e_st_eq_landau_n_rt_rp_ts s (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2))) (l_e_st_eq_landau_n_rt_rp_ts s a1) (l_e_st_eq_landau_n_rt_rp_iv4d_t4 a1 a2 r) (l_e_st_eq_landau_n_rt_rp_iv4d_t2 a1 a2 s))))). Time Defined. (* constant 3571 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t9 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b1) (l_e_st_eq_landau_n_rt_rp_ts a2 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_ts a2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b1) (l_e_st_eq_landau_n_rt_rp_ts a2 b2)) (l_e_st_eq_landau_n_rt_rp_iv4d_t5 a1 a2 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_iv4d_t6 b1 b2 a1 a2))))). Time Defined. (* constant 3572 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t10 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b2) (l_e_st_eq_landau_n_rt_rp_ts a2 b1)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_ts a2 (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b2) (l_e_st_eq_landau_n_rt_rp_ts a2 b1)) (l_e_st_eq_landau_n_rt_rp_iv4d_t5 a1 a2 (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_iv4d_t8 b1 b2 a1 a2))))). Time Defined. (* constant 3573 *) Definition l_e_st_eq_landau_n_rt_rp_td12 : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b1) (l_e_st_eq_landau_n_rt_rp_ts a2 b2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b2) (l_e_st_eq_landau_n_rt_rp_ts a2 b1))))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_issmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df b1 b2))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df a1 a2)) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df b1 b2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b1) (l_e_st_eq_landau_n_rt_rp_ts a2 b2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b2) (l_e_st_eq_landau_n_rt_rp_ts a2 b1)) (l_e_st_eq_landau_n_rt_rp_iv4d_t9 a1 a2 b1 b2) (l_e_st_eq_landau_n_rt_rp_iv4d_t10 a1 a2 b1 b2))))). Time Defined. (* constant 3574 *) Definition l_e_st_eq_landau_n_rt_rp_td1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r1)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_issmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r1)) (l_e_st_eq_landau_n_rt_rp_iv4d_t6 r1 r2 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_iv4d_t8 r1 r2 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a))))). Time Defined. (* constant 3575 *) Definition l_e_st_eq_landau_n_rt_rp_td2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_is l_e_st_eq_landau_n_rt_rp_dif (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_1a a))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_issmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_iv4d_t5 r1 r2 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_iv4d_t5 r1 r2 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a))))). Time Defined. (* constant 3576 *) Definition l_e_st_eq_landau_n_rt_rp_tdeq12a : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b1) (l_e_st_eq_landau_n_rt_rp_ts a2 b2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b2) (l_e_st_eq_landau_n_rt_rp_ts a2 b1))))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b1) (l_e_st_eq_landau_n_rt_rp_ts a2 b2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b2) (l_e_st_eq_landau_n_rt_rp_ts a2 b1))) (l_e_st_eq_landau_n_rt_rp_td12 a1 a2 b1 b2))))). Time Defined. (* constant 3577 *) Definition l_e_st_eq_landau_n_rt_rp_tdeq12b : (forall (a1:l_e_st_eq_landau_n_rt_cut), (forall (a2:l_e_st_eq_landau_n_rt_cut), (forall (b1:l_e_st_eq_landau_n_rt_cut), (forall (b2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b1) (l_e_st_eq_landau_n_rt_rp_ts a2 b2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b2) (l_e_st_eq_landau_n_rt_rp_ts a2 b1))) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)))))). exact (fun (a1:l_e_st_eq_landau_n_rt_cut) => (fun (a2:l_e_st_eq_landau_n_rt_cut) => (fun (b1:l_e_st_eq_landau_n_rt_cut) => (fun (b2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq2 (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df a1 a2) (l_e_st_eq_landau_n_rt_rp_df b1 b2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b1) (l_e_st_eq_landau_n_rt_rp_ts a2 b2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts a1 b2) (l_e_st_eq_landau_n_rt_rp_ts a2 b1))) (l_e_st_eq_landau_n_rt_rp_td12 a1 a2 b1 b2))))). Time Defined. (* constant 3578 *) Definition l_e_st_eq_landau_n_rt_rp_tdeq1a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r1)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r1))) (l_e_st_eq_landau_n_rt_rp_td1 a r1 r2)))). Time Defined. (* constant 3579 *) Definition l_e_st_eq_landau_n_rt_rp_tdeq1b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r1))) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_df r1 r2))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq2 (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_df r1 r2)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r1) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r2) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r1))) (l_e_st_eq_landau_n_rt_rp_td1 a r1 r2)))). Time Defined. (* constant 3580 *) Definition l_e_st_eq_landau_n_rt_rp_tdeq2a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_1a a))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq1 (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_1a a)))) (l_e_st_eq_landau_n_rt_rp_td2 a r1 r2)))). Time Defined. (* constant 3581 *) Definition l_e_st_eq_landau_n_rt_rp_tdeq2b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_1a a)))) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df r1 r2) a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_refeq2 (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_df r1 r2) a) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r1 (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts r2 (l_e_st_eq_landau_n_rt_rp_1a a)))) (l_e_st_eq_landau_n_rt_rp_td2 a r1 r2)))). Time Defined. (* constant 3582 *) Definition l_e_st_eq_landau_n_rt_rp_4d194_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))). Time Defined. (* constant 3583 *) Definition l_e_st_eq_landau_n_rt_rp_4d194_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_comts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))))). Time Defined. (* constant 3584 *) Definition l_e_st_eq_landau_n_rt_rp_satzd194 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_4d194_t1 a b) (l_e_st_eq_landau_n_rt_rp_4d194_t2 a b))). Time Defined. (* constant 3585 *) Definition l_e_st_eq_landau_n_rt_rp_comtd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td b a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satzd194 a b)). Time Defined. (* constant 3586 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t11 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) r)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) r)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) r) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) r)) (l_e_st_eq_landau_n_rt_rp_distpt1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b) r) (l_e_st_eq_landau_n_rt_rp_ists1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a)) r e) (l_e_st_eq_landau_n_rt_rp_disttp1 (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2a a) r)))))). Time Defined. (* constant 3587 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)))) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_iv4d_t11 a b c e (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_iv4d_t11 b a c (l_e_st_eq_landau_n_rt_rp_symeq a b e) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))))))). Time Defined. (* constant 3588 *) Definition l_e_st_eq_landau_n_rt_rp_eqtd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_iv4d_t12 a b c e))))). Time Defined. (* constant 3589 *) Definition l_e_st_eq_landau_n_rt_rp_eqtd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_td c b) (l_e_st_eq_landau_n_rt_rp_comtd c a) (l_e_st_eq_landau_n_rt_rp_eqtd1 a b c e) (l_e_st_eq_landau_n_rt_rp_comtd b c))))). Time Defined. (* constant 3590 *) Definition l_e_st_eq_landau_n_rt_rp_eqtd12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_td b d) (l_e_st_eq_landau_n_rt_rp_eqtd1 a b c e) (l_e_st_eq_landau_n_rt_rp_eqtd2 c d b f))))))). Time Defined. (* constant 3591 *) Definition l_e_st_eq_landau_n_rt_rp_4d192_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ists1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b) z) (l_e_st_eq_landau_n_rt_rp_ists1 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_symis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) z))) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))))). Time Defined. (* constant 3592 *) Definition l_e_st_eq_landau_n_rt_rp_satzd192a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_zeroi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_4d192_t1 a b z)))). Time Defined. (* constant 3593 *) Definition l_e_st_eq_landau_n_rt_rp_satzd192b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_comtd b a) (l_e_st_eq_landau_n_rt_rp_satzd192a b a z)))). Time Defined. (* constant 3594 *) Definition l_e_st_eq_landau_n_rt_rp_td01 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_satzd192a a b z))). Time Defined. (* constant 3595 *) Definition l_e_st_eq_landau_n_rt_rp_td02 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_satzd192b a b z))). Time Defined. (* constant 3596 *) Definition l_e_st_eq_landau_n_rt_rp_satzd197a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) b) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)))) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_tdeq2a b (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)))) (l_e_st_eq_landau_n_rt_rp_m0deqb (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)))))). Time Defined. (* constant 3597 *) Definition l_e_st_eq_landau_n_rt_rp_satzd197b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d b) a) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td b a)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_comtd a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_satzd197a b a) (l_e_st_eq_landau_n_rt_rp_eqm0d (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_comtd b a)))). Time Defined. (* constant 3598 *) Definition l_e_st_eq_landau_n_rt_rp_satzd197c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) b) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq2 (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) b) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_satzd197a a b) (l_e_st_eq_landau_n_rt_rp_satzd197b a b))). Time Defined. (* constant 3599 *) Definition l_e_st_eq_landau_n_rt_rp_satzd197d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) b) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_satzd197c a b))). Time Defined. (* constant 3600 *) Definition l_e_st_eq_landau_n_rt_rp_satzd197e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_satzd197a a b))). Time Defined. (* constant 3601 *) Definition l_e_st_eq_landau_n_rt_rp_satzd197f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_satzd197b a b))). Time Defined. (* constant 3602 *) Definition l_e_st_eq_landau_n_rt_rp_satzd198 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_td a b))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b))) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_satzd197c a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_eqtd2 (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b)) b a (l_e_st_eq_landau_n_rt_rp_satzd177 b)))). Time Defined. (* constant 3603 *) Definition l_e_st_eq_landau_n_rt_rp_satzd198a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_satzd198 a b))). Time Defined. (* constant 3604 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t13 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_distpt2 r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ists2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_1b a b) r (l_e_st_eq_landau_n_rt_rp_satz140e (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q))))))). Time Defined. (* constant 3605 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t14 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b)) s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b)) s))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b)) s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b))) s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b)) s) (l_e_st_eq_landau_n_rt_rp_asspl2 (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b)) s) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b)) s (l_e_st_eq_landau_n_rt_rp_iv4d_t13 a b p q r)))))))). Time Defined. (* constant 3606 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t15 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b)) s)) (l_e_st_eq_landau_n_rt_rp_pl s (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_2b a b)) s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b)) s) (l_e_st_eq_landau_n_rt_rp_pl s (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_iv4d_t14 a b p q r s) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_1b a b)) s))))))). Time Defined. (* constant 3607 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t16 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_satz135h (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_satz145a (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q) p))))). Time Defined. (* constant 3608 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t17 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) q)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_iv4d_t14 a b p q (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_iv4d_t15 a b p q (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_iv4d_t16 a b p q))))). Time Defined. (* constant 3609 *) Definition l_e_st_eq_landau_n_rt_rp_ptdpp : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_td a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_posdi (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_iv4d_t17 a b p q))))). Time Defined. (* constant 3610 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t18 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_satzd197b a b) (l_e_st_eq_landau_n_rt_rp_ptdpp a (l_e_st_eq_landau_n_rt_rp_m0d b) p (l_e_st_eq_landau_n_rt_rp_satzd176c b n)))))). Time Defined. (* constant 3611 *) Definition l_e_st_eq_landau_n_rt_rp_ntdpn : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_td a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_satzd176f (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_iv4d_t18 a b p n))))). Time Defined. (* constant 3612 *) Definition l_e_st_eq_landau_n_rt_rp_ntdnp : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (p:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_td a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_eqnegd (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_comtd b a) (l_e_st_eq_landau_n_rt_rp_ntdpn b a p n))))). Time Defined. (* constant 3613 *) Definition l_e_st_eq_landau_n_rt_rp_ptdnn : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (o:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_td a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_satzd198 a b) (l_e_st_eq_landau_n_rt_rp_ptdpp (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_satzd176c a n) (l_e_st_eq_landau_n_rt_rp_satzd176c b o)))))). Time Defined. (* constant 3614 *) Definition l_e_st_eq_landau_n_rt_rp_4d192_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_pnot0d (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_ptdpp a b p q))))))). Time Defined. (* constant 3615 *) Definition l_e_st_eq_landau_n_rt_rp_4d192_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (m:l_e_st_eq_landau_n_rt_rp_negd b), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_nnot0d (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_ntdpn a b p m))))))). Time Defined. (* constant 3616 *) Definition l_e_st_eq_landau_n_rt_rp_4d192_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_rappd b (l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) (fun (t:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_4d192_t2 a b n o p t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_zero b) (l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) o) (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_4d192_t3 a b n o p t)))))). Time Defined. (* constant 3617 *) Definition l_e_st_eq_landau_n_rt_rp_4d192_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), (forall (p:l_e_st_eq_landau_n_rt_rp_posd b), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_nnot0d (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_ntdnp a b m p))))))). Time Defined. (* constant 3618 *) Definition l_e_st_eq_landau_n_rt_rp_4d192_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), (forall (l:l_e_st_eq_landau_n_rt_rp_negd b), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (l:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_pnot0d (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_ptdnn a b m l))))))). Time Defined. (* constant 3619 *) Definition l_e_st_eq_landau_n_rt_rp_4d192_t7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_rappd b (l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) (fun (t:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_4d192_t5 a b n o m t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_zero b) (l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) o) (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_4d192_t6 a b n o m t)))))). Time Defined. (* constant 3620 *) Definition l_e_st_eq_landau_n_rt_rp_satzd192d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_e_st_eq_landau_n_rt_rp_rappd a (l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_4d192_t4 a b n o t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) n) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_4d192_t7 a b n o t))))). Time Defined. (* constant 3621 *) Definition l_e_st_eq_landau_n_rt_rp_4d192_t8 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_zero b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_et (l_e_st_eq_landau_n_rt_rp_zero b) (l_imp_th3 (l_not (l_e_st_eq_landau_n_rt_rp_zero b)) (l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) (l_weli (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)) z) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_e_st_eq_landau_n_rt_rp_satzd192d a b n t)))))). Time Defined. (* constant 3622 *) Definition l_e_st_eq_landau_n_rt_rp_satzd192c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)), l_or (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_zero b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)) => (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_4d192_t8 a b z t)))). Time Defined. (* constant 3623 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_treq2 (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_absnnd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_pnotnd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_ptdpp a b p q))) (l_e_st_eq_landau_n_rt_rp_eqtd12 (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absd b) b (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a p)) (l_e_st_eq_landau_n_rt_rp_absnnd b (l_e_st_eq_landau_n_rt_rp_pnotnd b q))))))). Time Defined. (* constant 3624 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_absnd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_ntdpn a b p n)) (l_e_st_eq_landau_n_rt_rp_satzd197f a b))))). Time Defined. (* constant 3625 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_treq2 (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_4d193_t2 a b p n) (l_e_st_eq_landau_n_rt_rp_eqtd12 (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a p)) (l_e_st_eq_landau_n_rt_rp_absnd b n)))))). Time Defined. (* constant 3626 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_zeroeq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_satzd166f (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td01 a b z)) (l_e_st_eq_landau_n_rt_rp_td01 (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_satzd166f a z))))). Time Defined. (* constant 3627 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_zeroeq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_satzd166f (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td02 a b z)) (l_e_st_eq_landau_n_rt_rp_td02 (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_satzd166f b z))))). Time Defined. (* constant 3628 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (p:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td b a)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_eqabsd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_comtd a b)) (l_e_st_eq_landau_n_rt_rp_4d193_t3 b a p n) (l_e_st_eq_landau_n_rt_rp_comtd (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_absd a)))))). Time Defined. (* constant 3629 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (o:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_eqtd12 (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_absnd a n) (l_e_st_eq_landau_n_rt_rp_absnd b o)) (l_e_st_eq_landau_n_rt_rp_satzd198 a b))))). Time Defined. (* constant 3630 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t8 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (o:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_treq2 (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_absnnd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_pnotnd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_ptdnn a b n o))) (l_e_st_eq_landau_n_rt_rp_4d193_t7 a b n o))))). Time Defined. (* constant 3631 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t9 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_rappd b (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) (fun (t:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_4d193_t1 a b p t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_4d193_t5 a b t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_4d193_t3 a b p t)))). Time Defined. (* constant 3632 *) Definition l_e_st_eq_landau_n_rt_rp_4d193_t10 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_rappd b (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) (fun (t:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_4d193_t6 a b n t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_4d193_t5 a b t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_4d193_t8 a b n t)))). Time Defined. (* constant 3633 *) Definition l_e_st_eq_landau_n_rt_rp_satzd193 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_rappd a (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_4d193_t9 a b t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_4d193_t4 a b t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_4d193_t10 a b t))). Time Defined. (* constant 3634 *) Definition l_e_st_eq_landau_n_rt_rp_satzd103a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_satzd193 a b))). Time Defined. (* constant 3635 *) Definition l_e_st_eq_landau_n_rt_rp_1df : l_e_st_eq_landau_n_rt_rp_dif. exact (l_e_st_eq_landau_n_rt_rp_pdofrp l_e_st_eq_landau_n_rt_rp_1rp). Time Defined. (* constant 3636 *) Definition l_e_st_eq_landau_n_rt_rp_4d195_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_a2isapa (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_satz151 (l_e_st_eq_landau_n_rt_rp_2a a))))). Time Defined. (* constant 3637 *) Definition l_e_st_eq_landau_n_rt_rp_4d195_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2a a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_asspl2 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_satz151 (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_a2isapa (l_e_st_eq_landau_n_rt_rp_2a a)))). Time Defined. (* constant 3638 *) Definition l_e_st_eq_landau_n_rt_rp_4d195_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1a a)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_4d195_t1 a) (l_e_st_eq_landau_n_rt_rp_4d195_t2 a)). Time Defined. (* constant 3639 *) Definition l_e_st_eq_landau_n_rt_rp_satzd195 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a l_e_st_eq_landau_n_rt_rp_1df) a). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td a l_e_st_eq_landau_n_rt_rp_1df) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)))) a (l_e_st_eq_landau_n_rt_rp_tdeq1a a (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_eqi2 a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_4d195_t3 a))). Time Defined. (* constant 3640 *) Definition l_e_st_eq_landau_n_rt_rp_satzd195a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_td a l_e_st_eq_landau_n_rt_rp_1df)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td a l_e_st_eq_landau_n_rt_rp_1df) a (l_e_st_eq_landau_n_rt_rp_satzd195 a)). Time Defined. (* constant 3641 *) Definition l_e_st_eq_landau_n_rt_rp_satzd195b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td l_e_st_eq_landau_n_rt_rp_1df a) a). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td l_e_st_eq_landau_n_rt_rp_1df a) (l_e_st_eq_landau_n_rt_rp_td a l_e_st_eq_landau_n_rt_rp_1df) a (l_e_st_eq_landau_n_rt_rp_comtd l_e_st_eq_landau_n_rt_rp_1df a) (l_e_st_eq_landau_n_rt_rp_satzd195 a)). Time Defined. (* constant 3642 *) Definition l_e_st_eq_landau_n_rt_rp_satzd195c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_td l_e_st_eq_landau_n_rt_rp_1df a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td l_e_st_eq_landau_n_rt_rp_1df a) a (l_e_st_eq_landau_n_rt_rp_satzd195b a)). Time Defined. (* constant 3643 *) Definition l_e_st_eq_landau_n_rt_rp_satzd196a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_eqtd12 (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_absd b) b (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a p)) (l_e_st_eq_landau_n_rt_rp_absnnd b (l_e_st_eq_landau_n_rt_rp_pnotnd b q))))))). Time Defined. (* constant 3644 *) Definition l_e_st_eq_landau_n_rt_rp_satzd196b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (o:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (o:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_treq2 (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_satzd198a a b) (l_e_st_eq_landau_n_rt_rp_eqtd12 (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_absnd a n) (l_e_st_eq_landau_n_rt_rp_absnd b o)))))). Time Defined. (* constant 3645 *) Definition l_e_st_eq_landau_n_rt_rp_satzd196c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_treq1 (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_absd b))) (l_e_st_eq_landau_n_rt_rp_eqtd12 (l_e_st_eq_landau_n_rt_rp_absd a) a (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_absd b)) b (l_e_st_eq_landau_n_rt_rp_absnnd a (l_e_st_eq_landau_n_rt_rp_pnotnd a p)) (l_e_st_eq_landau_n_rt_rp_satzd177b (l_e_st_eq_landau_n_rt_rp_absd b) b (l_e_st_eq_landau_n_rt_rp_absnd b n))) (l_e_st_eq_landau_n_rt_rp_satzd197b (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))))). Time Defined. (* constant 3646 *) Definition l_e_st_eq_landau_n_rt_rp_satzd196d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (p:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_absd a))) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) (l_e_st_eq_landau_n_rt_rp_comtd a b) (l_e_st_eq_landau_n_rt_rp_satzd196c b a p n) (l_e_st_eq_landau_n_rt_rp_eqm0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_comtd (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_absd a))))))). Time Defined. (* constant 3647 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_p1p2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_posd b))). Time Defined. (* constant 3648 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_p1n2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd b))). Time Defined. (* constant 3649 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_n1p2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_posd b))). Time Defined. (* constant 3650 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_n1n2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_negd b))). Time Defined. (* constant 3651 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_e_st_eq_landau_n_rt_rp_ptdpp (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_satzd166e a n) (l_e_st_eq_landau_n_rt_rp_satzd166e b o))))). Time Defined. (* constant 3652 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), l_not (l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_td a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => l_e_st_eq_landau_n_rt_rp_pnotnd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) e) (l_e_st_eq_landau_n_rt_rp_4d196_t1 a b n o))))))). Time Defined. (* constant 3653 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_not (l_e_st_eq_landau_n_rt_rp_negd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_4d196_t2 a b n o e) (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_ntdpn a b p t))))))). Time Defined. (* constant 3654 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_posd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_axrdo b) (l_e_st_eq_landau_n_rt_rp_4d196_t3 a b n o e p) o)))))). Time Defined. (* constant 3655 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_4d196_p1p2 a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_andi (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_posd b) p (l_e_st_eq_landau_n_rt_rp_4d196_t4 a b n o e p))))))). Time Defined. (* constant 3656 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_or (l_e_st_eq_landau_n_rt_rp_4d196_p1p2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1n2 a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_ori1 (l_e_st_eq_landau_n_rt_rp_4d196_p1p2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1n2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_t5 a b n o e p))))))). Time Defined. (* constant 3657 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_not (l_e_st_eq_landau_n_rt_rp_posd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_4d196_t2 a b n o e) (fun (t:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_ntdnp a b m t))))))). Time Defined. (* constant 3658 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t8 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_negd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_or3e3 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_axrdo b) o (l_e_st_eq_landau_n_rt_rp_4d196_t7 a b n o e m))))))). Time Defined. (* constant 3659 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t9 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_4d196_n1n2 a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_andi (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_negd b) m (l_e_st_eq_landau_n_rt_rp_4d196_t8 a b n o e m))))))). Time Defined. (* constant 3660 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t10 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_or (l_e_st_eq_landau_n_rt_rp_4d196_p1p2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1n2 a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_ori2 (l_e_st_eq_landau_n_rt_rp_4d196_p1p2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1n2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_t9 a b n o e m))))))). Time Defined. (* constant 3661 *) Definition l_e_st_eq_landau_n_rt_rp_satzd196e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))), l_or (l_and (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_posd b)) (l_and (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_negd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) => l_e_st_eq_landau_n_rt_rp_rappd a (l_or (l_e_st_eq_landau_n_rt_rp_4d196_p1p2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1n2 a b)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_4d196_t6 a b n o e t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_or (l_e_st_eq_landau_n_rt_rp_4d196_p1p2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1n2 a b)) n) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_4d196_t10 a b n o e t)))))). Time Defined. (* constant 3662 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t11 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_e_st_eq_landau_n_rt_rp_satzd176a (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_4d196_t1 a b n o))))). Time Defined. (* constant 3663 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), l_not (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_td a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => l_e_st_eq_landau_n_rt_rp_nnotpd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_eqnegd (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b))) e) (l_e_st_eq_landau_n_rt_rp_4d196_t11 a b n o))))))). Time Defined. (* constant 3664 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t13 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_not (l_e_st_eq_landau_n_rt_rp_posd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_4d196_t12 a b n o e) (fun (t:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_ptdpp a b p t))))))). Time Defined. (* constant 3665 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t14 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_negd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_or3e3 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_axrdo b) o (l_e_st_eq_landau_n_rt_rp_4d196_t13 a b n o e p))))))). Time Defined. (* constant 3666 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t15 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_4d196_p1n2 a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_andi (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd b) p (l_e_st_eq_landau_n_rt_rp_4d196_t14 a b n o e p))))))). Time Defined. (* constant 3667 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t16 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_or (l_e_st_eq_landau_n_rt_rp_4d196_p1n2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1p2 a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_ori1 (l_e_st_eq_landau_n_rt_rp_4d196_p1n2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1p2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_t15 a b n o e p))))))). Time Defined. (* constant 3668 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t17 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_not (l_e_st_eq_landau_n_rt_rp_negd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_4d196_t12 a b n o e) (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_ptdnn a b m t))))))). Time Defined. (* constant 3669 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t18 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_posd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_axrdo b) (l_e_st_eq_landau_n_rt_rp_4d196_t17 a b n o e m) o)))))). Time Defined. (* constant 3670 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t19 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_4d196_n1p2 a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_andi (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_posd b) m (l_e_st_eq_landau_n_rt_rp_4d196_t18 a b n o e m))))))). Time Defined. (* constant 3671 *) Definition l_e_st_eq_landau_n_rt_rp_4d196_t20 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), (forall (m:l_e_st_eq_landau_n_rt_rp_negd a), l_or (l_e_st_eq_landau_n_rt_rp_4d196_p1n2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1p2 a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => (fun (m:l_e_st_eq_landau_n_rt_rp_negd a) => l_ori2 (l_e_st_eq_landau_n_rt_rp_4d196_p1n2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1p2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_t19 a b n o e m))))))). Time Defined. (* constant 3672 *) Definition l_e_st_eq_landau_n_rt_rp_satzd196f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))), l_or (l_and (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd b)) (l_and (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_posd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)))) => l_e_st_eq_landau_n_rt_rp_rappd a (l_or (l_e_st_eq_landau_n_rt_rp_4d196_p1n2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1p2 a b)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_4d196_t16 a b n o e t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_or (l_e_st_eq_landau_n_rt_rp_4d196_p1n2 a b) (l_e_st_eq_landau_n_rt_rp_4d196_n1p2 a b)) n) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_4d196_t20 a b n o e t)))))). Time Defined. (* constant 3673 *) Definition l_e_st_eq_landau_n_rt_rp_4d199_t1 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (t:l_e_st_eq_landau_n_rt_cut), (forall (u:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts q s)) t) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts r t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t))))))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_cut) => (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts q s)) t) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts p r) t) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts q s) t)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts r t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t))) (l_e_st_eq_landau_n_rt_rp_disttp1 (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts q s) t) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts p r) t) (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts r t)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ts q s) t) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t)) (l_e_st_eq_landau_n_rt_rp_assts1 p r t) (l_e_st_eq_landau_n_rt_rp_assts1 q s t)))))))). Time Defined. (* constant 3674 *) Definition l_e_st_eq_landau_n_rt_rp_4d199_t2 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (t:l_e_st_eq_landau_n_rt_cut), (forall (u:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u))) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r u) (l_e_st_eq_landau_n_rt_rp_ts s t))))))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_cut) => (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t))) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r u) (l_e_st_eq_landau_n_rt_rp_ts s t))) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u))) (l_e_st_eq_landau_n_rt_rp_distpt2 q (l_e_st_eq_landau_n_rt_rp_ts r u) (l_e_st_eq_landau_n_rt_rp_ts s t)))))))). Time Defined. (* constant 3675 *) Definition l_e_st_eq_landau_n_rt_rp_4d199_t3 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (t:l_e_st_eq_landau_n_rt_cut), (forall (u:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts q s)) t) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p s) (l_e_st_eq_landau_n_rt_rp_ts q r)) u)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r t) (l_e_st_eq_landau_n_rt_rp_ts s u))) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r u) (l_e_st_eq_landau_n_rt_rp_ts s t)))))))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_cut) => (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts q s)) t) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p s) (l_e_st_eq_landau_n_rt_rp_ts q r)) u)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts r t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts s u)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts r t)) (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts s u))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r t) (l_e_st_eq_landau_n_rt_rp_ts s u))) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r u) (l_e_st_eq_landau_n_rt_rp_ts s t)))) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts q s)) t) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts r t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p s) (l_e_st_eq_landau_n_rt_rp_ts q r)) u) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts s u)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u))) (l_e_st_eq_landau_n_rt_rp_4d199_t1 p q r s t u) (l_e_st_eq_landau_n_rt_rp_4d199_t1 p q s r u t)) (l_e_st_eq_landau_n_rt_rp_4pl23 (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts r t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t)) (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts s u)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u))) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts r t)) (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_ts s u))) (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r t) (l_e_st_eq_landau_n_rt_rp_ts s u))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts s t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_ts r u))) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r u) (l_e_st_eq_landau_n_rt_rp_ts s t))) (l_e_st_eq_landau_n_rt_rp_distpt2 p (l_e_st_eq_landau_n_rt_rp_ts r t) (l_e_st_eq_landau_n_rt_rp_ts s u)) (l_e_st_eq_landau_n_rt_rp_4d199_t2 p q r s t u)))))))). Time Defined. (* constant 3676 *) Definition l_e_st_eq_landau_n_rt_rp_satzd199 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_td a b) c) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_td b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_td a b) c) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_1c a b c)))) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)))))) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_tdeq2a c (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b)))) (l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))))) (l_e_st_eq_landau_n_rt_rp_4d199_t3 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_4d199_t3 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_tdeq1b a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))))))). Time Defined. (* constant 3677 *) Definition l_e_st_eq_landau_n_rt_rp_asstd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_td a b) c) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_td b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satzd199 a b c))). Time Defined. (* constant 3678 *) Definition l_e_st_eq_landau_n_rt_rp_asstd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_td a b) c)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_td a b) c) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_satzd199 a b c)))). Time Defined. (* constant 3679 *) Definition l_e_st_eq_landau_n_rt_rp_4d201_t1 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (t:l_e_st_eq_landau_n_rt_cut), (forall (u:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_pl r t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl s u))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts q s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p t) (l_e_st_eq_landau_n_rt_rp_ts q u))))))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_cut) => (fun (u:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_pl r t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl s u))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts p t)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q s) (l_e_st_eq_landau_n_rt_rp_ts q u))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts q s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p t) (l_e_st_eq_landau_n_rt_rp_ts q u))) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts p (l_e_st_eq_landau_n_rt_rp_pl r t)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts p t)) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl s u)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q s) (l_e_st_eq_landau_n_rt_rp_ts q u)) (l_e_st_eq_landau_n_rt_rp_disttp2 p r t) (l_e_st_eq_landau_n_rt_rp_disttp2 q s u)) (l_e_st_eq_landau_n_rt_rp_4pl23 (l_e_st_eq_landau_n_rt_rp_ts p r) (l_e_st_eq_landau_n_rt_rp_ts p t) (l_e_st_eq_landau_n_rt_rp_ts q s) (l_e_st_eq_landau_n_rt_rp_ts q u)))))))). Time Defined. (* constant 3680 *) Definition l_e_st_eq_landau_n_rt_rp_satzd201 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))))) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c))))) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_tdeq1a a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)))) (l_e_st_eq_landau_n_rt_rp_4d201_t1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1c a b c) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_4d201_t1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b) (l_e_st_eq_landau_n_rt_rp_1b a b) (l_e_st_eq_landau_n_rt_rp_2c a b c) (l_e_st_eq_landau_n_rt_rp_1c a b c))) (l_e_st_eq_landau_n_rt_rp_pdeq12b (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2b a b)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1b a b))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_1c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_2c a b c))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2c a b c)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1c a b c))))))). Time Defined. (* constant 3681 *) Definition l_e_st_eq_landau_n_rt_rp_disttpd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_td c (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_comtd (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_satzd201 c a b) (l_e_st_eq_landau_n_rt_rp_eqpd12 (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td c b) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_comtd c a) (l_e_st_eq_landau_n_rt_rp_comtd c b))))). Time Defined. (* constant 3682 *) Definition l_e_st_eq_landau_n_rt_rp_disttpd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satzd201 a b c))). Time Defined. (* constant 3683 *) Definition l_e_st_eq_landau_n_rt_rp_distptd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) c)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) c) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_disttpd1 a b c)))). Time Defined. (* constant 3684 *) Definition l_e_st_eq_landau_n_rt_rp_distptd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_pd b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_pd b c)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_disttpd2 a b c)))). Time Defined. (* constant 3685 *) Definition l_e_st_eq_landau_n_rt_rp_satzd202 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_md b c)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_md b c)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d c))) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_disttpd2 a b (l_e_st_eq_landau_n_rt_rp_m0d c)) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d c)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_satzd197b a c))))). Time Defined. (* constant 3686 *) Definition l_e_st_eq_landau_n_rt_rp_disttmd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_md a b) c) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_md a b) c) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d b) c)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_disttpd1 a (l_e_st_eq_landau_n_rt_rp_m0d b) c) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d b) c) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_satzd197a b c))))). Time Defined. (* constant 3687 *) Definition l_e_st_eq_landau_n_rt_rp_disttmd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_md b c)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satzd202 a b c))). Time Defined. (* constant 3688 *) Definition l_e_st_eq_landau_n_rt_rp_distmtd1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_md a b) c)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_md a b) c) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_disttmd1 a b c)))). Time Defined. (* constant 3689 *) Definition l_e_st_eq_landau_n_rt_rp_distmtd2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_md b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_md b c)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_disttmd2 a b c)))). Time Defined. (* constant 3690 *) Definition l_e_st_eq_landau_n_rt_rp_satzd200 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_md b c)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_td a c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satzd202 a b c))). Time Defined. (* constant 3691 *) Definition l_e_st_eq_landau_n_rt_rp_4d203_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => l_e_st_eq_landau_n_rt_rp_satzd182d a b m)))). Time Defined. (* constant 3692 *) Definition l_e_st_eq_landau_n_rt_rp_4d203_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd c), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd c) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_md a b) c) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_disttmd1 a b c) (l_e_st_eq_landau_n_rt_rp_ptdpp (l_e_st_eq_landau_n_rt_rp_md a b) c (l_e_st_eq_landau_n_rt_rp_4d203_t1 a b c m) p)))))). Time Defined. (* constant 3693 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd c), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd c) => l_e_st_eq_landau_n_rt_rp_satzd182a (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_4d203_t2 a b c m p)))))). Time Defined. (* constant 3694 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero c), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero c) => l_e_st_eq_landau_n_rt_rp_zeroeq (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_td02 a c z) (l_e_st_eq_landau_n_rt_rp_td02 b c z)))))). Time Defined. (* constant 3695 *) Definition l_e_st_eq_landau_n_rt_rp_4d203_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd c), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd c) => l_e_st_eq_landau_n_rt_rp_eqnegd (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_md a b) c) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)) (l_e_st_eq_landau_n_rt_rp_disttmd1 a b c) (l_e_st_eq_landau_n_rt_rp_ntdpn (l_e_st_eq_landau_n_rt_rp_md a b) c (l_e_st_eq_landau_n_rt_rp_4d203_t1 a b c m) n)))))). Time Defined. (* constant 3696 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203c : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd c), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd c) => l_e_st_eq_landau_n_rt_rp_satzd182c (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_4d203_t3 a b c m n)))))). Time Defined. (* constant 3697 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd c), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd c) => l_e_st_eq_landau_n_rt_rp_eqmored12 (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_td c b) (l_e_st_eq_landau_n_rt_rp_comtd a c) (l_e_st_eq_landau_n_rt_rp_comtd b c) (l_e_st_eq_landau_n_rt_rp_satzd203a a b c m p)))))). Time Defined. (* constant 3698 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203e : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero c), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero c) => l_e_st_eq_landau_n_rt_rp_zeroeq (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b) (l_e_st_eq_landau_n_rt_rp_td01 c a z) (l_e_st_eq_landau_n_rt_rp_td01 c b z)))))). Time Defined. (* constant 3699 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203f : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd c), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd c) => l_e_st_eq_landau_n_rt_rp_eqlessd12 (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_td c b) (l_e_st_eq_landau_n_rt_rp_comtd a c) (l_e_st_eq_landau_n_rt_rp_comtd b c) (l_e_st_eq_landau_n_rt_rp_satzd203c a b c m n)))))). Time Defined. (* constant 3700 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203g : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd c), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd c) => l_e_st_eq_landau_n_rt_rp_lemmad5 (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_satzd203a b a c (l_e_st_eq_landau_n_rt_rp_lemmad6 a b l) p)))))). Time Defined. (* constant 3701 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203h : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero c), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero c) => l_e_st_eq_landau_n_rt_rp_zeroeq (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_td02 a c z) (l_e_st_eq_landau_n_rt_rp_td02 b c z)))))). Time Defined. (* constant 3702 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203j : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd c), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd c) => l_e_st_eq_landau_n_rt_rp_lemmad6 (l_e_st_eq_landau_n_rt_rp_td b c) (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_satzd203c b a c (l_e_st_eq_landau_n_rt_rp_lemmad6 a b l) n)))))). Time Defined. (* constant 3703 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203k : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (p:l_e_st_eq_landau_n_rt_rp_posd c), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd c) => l_e_st_eq_landau_n_rt_rp_lemmad5 (l_e_st_eq_landau_n_rt_rp_td c b) (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_satzd203d b a c (l_e_st_eq_landau_n_rt_rp_lemmad6 a b l) p)))))). Time Defined. (* constant 3704 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203l : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero c), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero c) => l_e_st_eq_landau_n_rt_rp_zeroeq (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b) (l_e_st_eq_landau_n_rt_rp_td01 c a z) (l_e_st_eq_landau_n_rt_rp_td01 c b z)))))). Time Defined. (* constant 3705 *) Definition l_e_st_eq_landau_n_rt_rp_satzd203m : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a b), (forall (n:l_e_st_eq_landau_n_rt_rp_negd c), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_td c b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a b) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd c) => l_e_st_eq_landau_n_rt_rp_lemmad6 (l_e_st_eq_landau_n_rt_rp_td c b) (l_e_st_eq_landau_n_rt_rp_td c a) (l_e_st_eq_landau_n_rt_rp_satzd203f b a c (l_e_st_eq_landau_n_rt_rp_lemmad6 a b l) n)))))). Time Defined. (* constant 3706 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t19 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) q))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) (l_e_st_eq_landau_n_rt_rp_ts q l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) q) (l_e_st_eq_landau_n_rt_rp_disttp2 q p l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_ts q l_e_st_eq_landau_n_rt_rp_1rp) q (l_e_st_eq_landau_n_rt_rp_ts q p) (l_e_st_eq_landau_n_rt_rp_satz151 q)))). Time Defined. (* constant 3707 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t20 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) (l_e_st_eq_landau_n_rt_rp_pl q r))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) q) r) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) (l_e_st_eq_landau_n_rt_rp_pl q r)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) q) (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_1rp) r (l_e_st_eq_landau_n_rt_rp_iv4d_t19 p q) (l_e_st_eq_landau_n_rt_rp_satz151 r)) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_ts q p) q r)))). Time Defined. (* constant 3708 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t21 : (forall (p:l_e_st_eq_landau_n_rt_cut), (forall (q:l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) (l_e_st_eq_landau_n_rt_rp_pl r q))))). exact (fun (p:l_e_st_eq_landau_n_rt_cut) => (fun (q:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) (l_e_st_eq_landau_n_rt_rp_pl q r)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts q p) (l_e_st_eq_landau_n_rt_rp_pl r q)) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts q (l_e_st_eq_landau_n_rt_rp_pl p l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_iv4d_t20 p q r) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl q r) (l_e_st_eq_landau_n_rt_rp_pl r q) (l_e_st_eq_landau_n_rt_rp_ts q p) (l_e_st_eq_landau_n_rt_rp_compl q r))))). Time Defined. (* constant 3709 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_arp : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_cut)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_rpofpd a p)). Time Defined. (* constant 3710 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_arpi : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_cut)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_ov l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_iv4d_arp a p))). Time Defined. (* constant 3711 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_ai : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_dif)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p))). Time Defined. (* constant 3712 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t22 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_iv4d_ai a p)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_iv4d_ai a p)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p) l_e_st_eq_landau_n_rt_rp_1rp)))) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)))) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p))) (l_e_st_eq_landau_n_rt_rp_tdeq1a a (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_eqsmsd (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p) l_e_st_eq_landau_n_rt_rp_1rp))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a))) (l_e_st_eq_landau_n_rt_rp_iv4d_t20 (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p) (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a)) (l_e_st_eq_landau_n_rt_rp_iv4d_t21 (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p) (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_1a a))) (l_e_st_eq_landau_n_rt_rp_lemmad2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a))))). Time Defined. (* constant 3713 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t23 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) l_e_st_eq_landau_n_rt_rp_1rp))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arp a p)) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv4d_arp a p) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ists1 (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arp a p)) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p) (l_e_st_eq_landau_n_rt_rp_satz140d (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_2a a) p)) (l_e_st_eq_landau_n_rt_rp_disttp1 (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arp a p) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv4d_arp a p) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_satz153c l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_iv4d_arp a p))))). Time Defined. (* constant 3714 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t24 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p))) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_iv4d_t23 a p)) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_compl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)))). Time Defined. (* constant 3715 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t25 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p))) l_e_st_eq_landau_n_rt_rp_1df)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_iv4d_t24 a p))). Time Defined. (* constant 3716 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t26 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_iv4d_ai a p)) l_e_st_eq_landau_n_rt_rp_1df)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_iv4d_ai a p)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_1a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_2a a) (l_e_st_eq_landau_n_rt_rp_iv4d_arpi a p))) l_e_st_eq_landau_n_rt_rp_1df (l_e_st_eq_landau_n_rt_rp_iv4d_t22 a p) (l_e_st_eq_landau_n_rt_rp_iv4d_t25 a p))). Time Defined. (* constant 3717 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t27 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df) (l_e_st_eq_landau_n_rt_rp_iv4d_ai a p) (l_e_st_eq_landau_n_rt_rp_iv4d_t26 a p))). Time Defined. (* constant 3718 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t28 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_satzd176c a n)). Time Defined. (* constant 3719 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t29 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) h) l_e_st_eq_landau_n_rt_rp_1df), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d h)) l_e_st_eq_landau_n_rt_rp_1df)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) h) l_e_st_eq_landau_n_rt_rp_1df) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_m0d h)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) h) l_e_st_eq_landau_n_rt_rp_1df (l_e_st_eq_landau_n_rt_rp_satzd197d a h) e)))). Time Defined. (* constant 3720 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t30 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) h) l_e_st_eq_landau_n_rt_rp_1df), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) h) l_e_st_eq_landau_n_rt_rp_1df) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df) (l_e_st_eq_landau_n_rt_rp_m0d h) (l_e_st_eq_landau_n_rt_rp_iv4d_t29 a n h e))))). Time Defined. (* constant 3721 *) Definition l_e_st_eq_landau_n_rt_rp_iv4d_t31 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) x) l_e_st_eq_landau_n_rt_rp_1df) (l_e_st_eq_landau_n_rt_rp_iv4d_t27 (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_iv4d_t28 a n)) (l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a) x) l_e_st_eq_landau_n_rt_rp_1df) => l_e_st_eq_landau_n_rt_rp_iv4d_t30 a n x t)))). Time Defined. (* constant 3722 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_rappd a (l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_iv4d_t27 a t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a x) l_e_st_eq_landau_n_rt_rp_1df)) n) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_iv4d_t31 a t))). Time Defined. (* constant 3723 *) Definition l_e_st_eq_landau_n_rt_rp_4d204_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) a), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b k) a), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) (l_e_st_eq_landau_n_rt_rp_td b k)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) a) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b k) a) => l_e_st_eq_landau_n_rt_rp_treq2 (l_e_st_eq_landau_n_rt_rp_td b h) (l_e_st_eq_landau_n_rt_rp_td b k) a e f))))))). Time Defined. (* constant 3724 *) Definition l_e_st_eq_landau_n_rt_rp_4d204_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) a), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b k) a), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td b (l_e_st_eq_landau_n_rt_rp_md h k))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) a) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b k) a) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td b h) (l_e_st_eq_landau_n_rt_rp_td b k)) (l_e_st_eq_landau_n_rt_rp_td b (l_e_st_eq_landau_n_rt_rp_md h k)) (l_e_st_eq_landau_n_rt_rp_distmtd2 b h k) (l_e_st_eq_landau_n_rt_rp_satzd182e (l_e_st_eq_landau_n_rt_rp_td b h) (l_e_st_eq_landau_n_rt_rp_td b k) (l_e_st_eq_landau_n_rt_rp_4d204_t1 a b n h k e f))))))))). Time Defined. (* constant 3725 *) Definition l_e_st_eq_landau_n_rt_rp_4d204_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) a), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b k) a), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md h k)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) a) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b k) a) => l_ore2 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md h k)) (l_e_st_eq_landau_n_rt_rp_satzd192c b (l_e_st_eq_landau_n_rt_rp_md h k) (l_e_st_eq_landau_n_rt_rp_4d204_t2 a b n h k e f)) n))))))). Time Defined. (* constant 3726 *) Definition l_e_st_eq_landau_n_rt_rp_satzd204b : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (k:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) a), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b k) a), l_e_st_eq_landau_n_rt_rp_eq h k))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (k:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) a) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b k) a) => l_e_st_eq_landau_n_rt_rp_satzd182b h k (l_e_st_eq_landau_n_rt_rp_4d204_t3 a b n h k e f)))))))). Time Defined. (* constant 3727 *) Definition l_e_st_eq_landau_n_rt_rp_4d204_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) l_e_st_eq_landau_n_rt_rp_1df), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b (l_e_st_eq_landau_n_rt_rp_td h a)) a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) l_e_st_eq_landau_n_rt_rp_1df) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td b (l_e_st_eq_landau_n_rt_rp_td h a)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_td b h) a) (l_e_st_eq_landau_n_rt_rp_td l_e_st_eq_landau_n_rt_rp_1df a) a (l_e_st_eq_landau_n_rt_rp_asstd2 b h a) (l_e_st_eq_landau_n_rt_rp_eqtd1 (l_e_st_eq_landau_n_rt_rp_td b h) l_e_st_eq_landau_n_rt_rp_1df a e) (l_e_st_eq_landau_n_rt_rp_satzd195b a)))))). Time Defined. (* constant 3728 *) Definition l_e_st_eq_landau_n_rt_rp_4d204_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), (forall (h:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) l_e_st_eq_landau_n_rt_rp_1df), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b x) a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => (fun (h:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b h) l_e_st_eq_landau_n_rt_rp_1df) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b x) a) (l_e_st_eq_landau_n_rt_rp_td h a) (l_e_st_eq_landau_n_rt_rp_4d204_t4 a b n h e)))))). Time Defined. (* constant 3729 *) Definition l_e_st_eq_landau_n_rt_rp_satzd204a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b x) a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero b)) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b x) l_e_st_eq_landau_n_rt_rp_1df) (l_e_st_eq_landau_n_rt_rp_lemmad7 b n) (l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b x) a)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b x) l_e_st_eq_landau_n_rt_rp_1df) => l_e_st_eq_landau_n_rt_rp_4d204_t5 a b n x t))))). Time Defined. (* constant 3730 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t1 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r s), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r s) => l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl s (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_asspl2 r l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_asspl2 s l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_satz134 r s (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) m)))). Time Defined. (* constant 3731 *) Definition l_e_st_eq_landau_n_rt_rp_morerpepd : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r s), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r s) => l_e_st_eq_landau_n_rt_rp_moredi12 (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_iv5d_t1 r s m)))). Time Defined. (* constant 3732 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t2 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) => l_e_st_eq_landau_n_rt_rp_morede12 (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp m))). Time Defined. (* constant 3733 *) Definition l_e_st_eq_landau_n_rt_rp_morerpipd : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)), l_e_st_eq_landau_n_rt_rp_more r s))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) => l_e_st_eq_landau_n_rt_rp_satz136a r s (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ismore12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl r (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl s (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_asspl1 r l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_asspl1 s l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_iv5d_t2 r s m))))). Time Defined. (* constant 3734 *) Definition l_e_st_eq_landau_n_rt_rp_lessrpepd : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r s), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r s) => l_e_st_eq_landau_n_rt_rp_lemmad5 (l_e_st_eq_landau_n_rt_rp_pdofrp s) (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_morerpepd s r (l_e_st_eq_landau_n_rt_rp_satz122 r s l))))). Time Defined. (* constant 3735 *) Definition l_e_st_eq_landau_n_rt_rp_lessrpipd : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)), l_e_st_eq_landau_n_rt_rp_less r s))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) => l_e_st_eq_landau_n_rt_rp_satz121 s r (l_e_st_eq_landau_n_rt_rp_morerpipd s r (l_e_st_eq_landau_n_rt_rp_lemmad6 (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s) l))))). Time Defined. (* constant 3736 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_i : l_e_st_eq_landau_n_rt_cut. exact l_e_st_eq_landau_n_rt_rp_1rp. Time Defined. (* constant 3737 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_2 : l_e_st_eq_landau_n_rt_cut. exact (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i). Time Defined. (* constant 3738 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_rp1 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_iv5d_i). Time Defined. (* constant 3739 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_sp1 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_iv5d_i)). Time Defined. (* constant 3740 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_rps : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pl r s)). Time Defined. (* constant 3741 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_rs : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ts r s)). Time Defined. (* constant 3742 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t3 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) l_e_st_eq_landau_n_rt_rp_iv5d_2))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) l_e_st_eq_landau_n_rt_rp_iv5d_2) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_4pl23 r l_e_st_eq_landau_n_rt_rp_iv5d_i s l_e_st_eq_landau_n_rt_rp_iv5d_i)) (l_e_st_eq_landau_n_rt_rp_3pl23 (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2 l_e_st_eq_landau_n_rt_rp_iv5d_i))). Time Defined. (* constant 3743 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t4 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) l_e_st_eq_landau_n_rt_rp_iv5d_2) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)) (l_e_st_eq_landau_n_rt_rp_pdeq12a (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) l_e_st_eq_landau_n_rt_rp_iv5d_2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_t3 r s)))). Time Defined. (* constant 3744 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad8 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_iv5d_t4 r s)). Time Defined. (* constant 3745 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t5 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) r))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_iv5d_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) r) (l_e_st_eq_landau_n_rt_rp_disttp2 r s l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_ts r l_e_st_eq_landau_n_rt_rp_iv5d_i) r (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_satz151 r)))). Time Defined. (* constant 3746 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t6 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)) l_e_st_eq_landau_n_rt_rp_iv5d_i))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tr4is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) r) s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_disttp1 r l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts r (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) r) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s) (l_e_st_eq_landau_n_rt_rp_iv5d_t5 r s) (l_e_st_eq_landau_n_rt_rp_satz151b (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s))) (l_e_st_eq_landau_n_rt_rp_asspl2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) r) s l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) r) s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) r s)))). Time Defined. (* constant 3747 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t7 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)) l_e_st_eq_landau_n_rt_rp_iv5d_i) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)) l_e_st_eq_landau_n_rt_rp_iv5d_2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_t6 r s) (l_e_st_eq_landau_n_rt_rp_satz151 l_e_st_eq_landau_n_rt_rp_iv5d_i)) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s)) l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2))). Time Defined. (* constant 3748 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t8 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s) (l_e_st_eq_landau_n_rt_rp_satz151 (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r)) (l_e_st_eq_landau_n_rt_rp_satz151b (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s))) (l_e_st_eq_landau_n_rt_rp_4pl23 r l_e_st_eq_landau_n_rt_rp_iv5d_i s l_e_st_eq_landau_n_rt_rp_iv5d_i))). Time Defined. (* constant 3749 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t9 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2)) l_e_st_eq_landau_n_rt_rp_iv5d_i))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2)) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_iv5d_t8 r s)) (l_e_st_eq_landau_n_rt_rp_3pl23 (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2)))). Time Defined. (* constant 3750 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t10 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i)) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i)) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2)) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rps r s) l_e_st_eq_landau_n_rt_rp_iv5d_2)) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_t7 r s)) (l_e_st_eq_landau_n_rt_rp_iv5d_t9 r s))). Time Defined. (* constant 3751 *) Definition l_e_st_eq_landau_n_rt_rp_iv5d_t11 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s)) (l_e_st_eq_landau_n_rt_rp_tdeq12a (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i l_e_st_eq_landau_n_rt_rp_iv5d_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_iv5d_rp1 r) l_e_st_eq_landau_n_rt_rp_iv5d_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_sp1 r s))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_iv5d_rs r s) l_e_st_eq_landau_n_rt_rp_iv5d_i) l_e_st_eq_landau_n_rt_rp_iv5d_i (l_e_st_eq_landau_n_rt_rp_iv5d_t10 r s)))). Time Defined. (* constant 3752 *) Definition l_e_st_eq_landau_n_rt_rp_lemmad9 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_iv5d_t11 r s)). Time Defined. (* constant 3753 *) Definition l_e_st_eq_landau_n_rt_rp_in : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => l_e_st_esti l_e_st_eq_landau_n_rt_cut r s0)). Time Defined. (* constant 3754 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_prop1 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), Prop))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_less x r), l_e_st_eq_landau_n_rt_rp_in x s0))))). Time Defined. (* constant 3755 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_prop2 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), Prop))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_more x r), l_e_st_eq_landau_n_rt_rp_in x t0))))). Time Defined. (* constant 3756 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_prop3 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (r:l_e_st_eq_landau_n_rt_cut), Prop))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (r:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_5p205_prop1 s0 t0 r) (l_e_st_eq_landau_n_rt_rp_5p205_prop2 s0 t0 r)))). Time Defined. (* constant 3757 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t1 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), l_e_st_eq_landau_n_rt_rp_5p205_prop2 s0 t0 r1)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => l_ande2 (l_e_st_eq_landau_n_rt_rp_5p205_prop1 s0 t0 r1) (l_e_st_eq_landau_n_rt_rp_5p205_prop2 s0 t0 r1) pr1)))))))))). Time Defined. (* constant 3758 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t2 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), l_e_st_eq_landau_n_rt_rp_5p205_prop1 s0 t0 r2)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => l_ande1 (l_e_st_eq_landau_n_rt_rp_5p205_prop1 s0 t0 r2) (l_e_st_eq_landau_n_rt_rp_5p205_prop2 s0 t0 r2) pr2)))))))))). Time Defined. (* constant 3759 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_rx : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), (forall (l:l_e_st_eq_landau_n_rt_rp_less r1 r2), (forall (x0:l_e_st_eq_landau_n_rt_rat), l_e_st_eq_landau_n_rt_cut)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r1 r2) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_rpofrt x0)))))))))))). Time Defined. (* constant 3760 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t3 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), (forall (l:l_e_st_eq_landau_n_rt_rp_less r1 r2), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less r1 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) r2), l_e_st_eq_landau_n_rt_rp_in (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) s0)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r1 r2) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less r1 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) r2) => l_e_st_eq_landau_n_rt_rp_5p205_t2 s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) l2)))))))))))))). Time Defined. (* constant 3761 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t4 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), (forall (l:l_e_st_eq_landau_n_rt_rp_less r1 r2), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less r1 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) r2), l_e_st_eq_landau_n_rt_rp_in (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) t0)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r1 r2) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less r1 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) r2) => l_e_st_eq_landau_n_rt_rp_5p205_t1 s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) (l_e_st_eq_landau_n_rt_rp_satz122 r1 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) l1))))))))))))))). Time Defined. (* constant 3762 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t5 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), (forall (l:l_e_st_eq_landau_n_rt_rp_less r1 r2), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (l1:l_e_st_eq_landau_n_rt_rp_less r1 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)), (forall (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) r2), l_con)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r1 r2) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (l1:l_e_st_eq_landau_n_rt_rp_less r1 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)) => (fun (l2:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) r2) => l_ec3e31 (l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)) (l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)) (l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)) (l_e_st_eq_landau_n_rt_rp_satz123b (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)) (p2 (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) (l_e_st_eq_landau_n_rt_rp_5p205_t3 s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0 l1 l2) (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0) (l_e_st_eq_landau_n_rt_rp_5p205_t4 s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0 l1 l2)) (l_e_refis l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_5p205_rx s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x0)))))))))))))))). Time Defined. (* constant 3763 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t6 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), (forall (l:l_e_st_eq_landau_n_rt_rp_less r1 r2), l_con))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r1 r2) => l_e_st_eq_landau_n_rt_rp_satz159app r1 r2 l l_con (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_rp_less r1 (l_e_st_eq_landau_n_rt_rp_rpofrt x)) => (fun (u:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_rpofrt x) r2) => l_e_st_eq_landau_n_rt_rp_5p205_t5 s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 l x t u)))))))))))))). Time Defined. (* constant 3764 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t7 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), l_not (l_e_st_eq_landau_n_rt_rp_less r1 r2))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => (fun (t:l_e_st_eq_landau_n_rt_rp_less r1 r2) => l_e_st_eq_landau_n_rt_rp_5p205_t6 s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2 t))))))))))). Time Defined. (* constant 3765 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t8 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), l_not (l_e_st_eq_landau_n_rt_rp_more r1 r2))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => (fun (t:l_e_st_eq_landau_n_rt_rp_more r1 r2) => l_e_st_eq_landau_n_rt_rp_5p205_t6 s0 t0 p0 p1a p1b p2 r2 r1 pr2 pr1 (l_e_st_eq_landau_n_rt_rp_satz121 r1 r2 t)))))))))))). Time Defined. (* constant 3766 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t9 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r1:l_e_st_eq_landau_n_rt_cut), (forall (r2:l_e_st_eq_landau_n_rt_cut), (forall (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1), (forall (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2), l_e_st_eq_landau_n_rt_rp_is r1 r2)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r1:l_e_st_eq_landau_n_rt_cut) => (fun (r2:l_e_st_eq_landau_n_rt_cut) => (fun (pr1:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r1) => (fun (pr2:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 r2) => l_or3e1 (l_e_st_eq_landau_n_rt_rp_is r1 r2) (l_e_st_eq_landau_n_rt_rp_more r1 r2) (l_e_st_eq_landau_n_rt_rp_less r1 r2) (l_e_st_eq_landau_n_rt_rp_satz123a r1 r2) (l_e_st_eq_landau_n_rt_rp_5p205_t8 s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2) (l_e_st_eq_landau_n_rt_rp_5p205_t7 s0 t0 p0 p1a p1b p2 r1 r2 pr1 pr2))))))))))). Time Defined. (* constant 3767 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t10 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_amone l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 x))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_cut) => (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 x) => (fun (u:l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 y) => l_e_st_eq_landau_n_rt_rp_5p205_t9 s0 t0 p0 p1a p1b p2 x y t u)))))))))). Time Defined. (* constant 3768 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_schnittprop : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_some (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0))))))))). Time Defined. (* constant 3769 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_schnittset : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_set l_e_st_eq_landau_n_rt_rat)))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_e_st_setof l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x))))))). Time Defined. (* constant 3770 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t11 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => l_andi (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0) i lx)))))))))). Time Defined. (* constant 3771 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t12 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x0)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => l_somei l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) r (l_e_st_eq_landau_n_rt_rp_5p205_t11 s0 t0 p0 p1a p1b p2 r i x0 lx))))))))))). Time Defined. (* constant 3772 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t13 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x) x0 (l_e_st_eq_landau_n_rt_rp_5p205_t12 s0 t0 p0 p1a p1b p2 r i x0 lx))))))))))). Time Defined. (* constant 3773 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t14 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r t0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (j:l_e_st_eq_landau_n_rt_rp_in s s0), l_e_st_eq_landau_n_rt_rp_more r s)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r t0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (j:l_e_st_eq_landau_n_rt_rp_in s s0) => l_e_st_eq_landau_n_rt_rp_satz122 s r (p2 s j r i))))))))))))). Time Defined. (* constant 3774 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t15 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r t0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (j:l_e_st_eq_landau_n_rt_rp_in s s0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) r)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r t0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (j:l_e_st_eq_landau_n_rt_rp_in s s0) => l_e_st_eq_landau_n_rt_rp_satz158b r x0 ux)))))))))))). Time Defined. (* constant 3775 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t16 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r t0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (j:l_e_st_eq_landau_n_rt_rp_in s s0), l_e_st_eq_landau_n_rt_rp_moreis (l_e_st_eq_landau_n_rt_rp_rpofrt x0) s)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r t0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (j:l_e_st_eq_landau_n_rt_rp_in s s0) => l_e_st_eq_landau_n_rt_rp_moreisi1 (l_e_st_eq_landau_n_rt_rp_rpofrt x0) s (l_e_st_eq_landau_n_rt_rp_satz127c (l_e_st_eq_landau_n_rt_rp_rpofrt x0) r s (l_e_st_eq_landau_n_rt_rp_5p205_t15 s0 t0 p0 p1a p1b p2 r i x0 ux s j) (l_e_st_eq_landau_n_rt_rp_5p205_t14 s0 t0 p0 p1a p1b p2 r i x0 ux s j)))))))))))))). Time Defined. (* constant 3776 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t17 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r t0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (j:l_e_st_eq_landau_n_rt_rp_in s s0), l_e_st_eq_landau_n_rt_urt s x0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r t0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (j:l_e_st_eq_landau_n_rt_rp_in s s0) => l_e_st_eq_landau_n_rt_rp_satz158d s x0 (l_e_st_eq_landau_n_rt_rp_5p205_t16 s0 t0 p0 p1a p1b p2 r i x0 ux s j))))))))))))). Time Defined. (* constant 3777 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t18 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r t0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (s:l_e_st_eq_landau_n_rt_cut), l_not (l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r t0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_weli (l_ec (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)) (fun (t:l_e_st_eq_landau_n_rt_rp_in s s0) => l_e_st_eq_landau_n_rt_rp_5p205_t17 s0 t0 p0 p1a p1b p2 r i x0 ux s t)))))))))))). Time Defined. (* constant 3778 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t19 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r t0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), l_not (l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x0))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r t0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => l_some_th5 l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) (fun (y:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_5p205_t18 s0 t0 p0 p1a p1b p2 r i x0 ux y))))))))))). Time Defined. (* constant 3779 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t20 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r t0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), l_not (l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r t0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => l_imp_th3 (l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x0) (l_e_st_eq_landau_n_rt_rp_5p205_t19 s0 t0 p0 p1a p1b p2 r i x0 ux) (fun (t:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x) x0 t))))))))))). Time Defined. (* constant 3780 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t21 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x0)))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x) x0 i)))))))). Time Defined. (* constant 3781 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t22 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), l_e_st_eq_landau_n_rt_rp_in r s0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0) a)))))))))))). Time Defined. (* constant 3782 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t23 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), l_e_st_eq_landau_n_rt_lrt r x0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0) a)))))))))))). Time Defined. (* constant 3783 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t24 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), l_e_st_eq_landau_n_rt_lrt r y0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => l_e_st_eq_landau_n_rt_rp_satz120 r x0 (l_e_st_eq_landau_n_rt_rp_5p205_t23 s0 t0 p0 p1a p1b p2 x0 i y0 l r a) y0 l)))))))))))). Time Defined. (* constant 3784 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t25 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r y0))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => l_andi (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r y0) (l_e_st_eq_landau_n_rt_rp_5p205_t22 s0 t0 p0 p1a p1b p2 x0 i y0 l r a) (l_e_st_eq_landau_n_rt_rp_5p205_t24 s0 t0 p0 p1a p1b p2 x0 i y0 l r a))))))))))))). Time Defined. (* constant 3785 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t26 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 y0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => l_somei l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y y0)) r (l_e_st_eq_landau_n_rt_rp_5p205_t25 s0 t0 p0 p1a p1b p2 x0 i y0 l r a))))))))))))). Time Defined. (* constant 3786 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t27 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 y0)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_someapp l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) (l_e_st_eq_landau_n_rt_rp_5p205_t21 s0 t0 p0 p1a p1b p2 x0 i) (l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 y0) (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (r:l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) => l_e_st_eq_landau_n_rt_rp_5p205_t26 s0 t0 p0 p1a p1b p2 x0 i y0 l y r)))))))))))). Time Defined. (* constant 3787 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t28 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (l:l_e_st_eq_landau_n_rt_less y0 x0), l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (l:l_e_st_eq_landau_n_rt_less y0 x0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x) y0 (l_e_st_eq_landau_n_rt_rp_5p205_t27 s0 t0 p0 p1a p1b p2 x0 i y0 l))))))))))). Time Defined. (* constant 3788 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t29 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), l_e_st_eq_landau_n_rt_rp_in r s0)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0) a)))))))))). Time Defined. (* constant 3789 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t30 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), l_e_st_eq_landau_n_rt_lrt r x0)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0) a)))))))))). Time Defined. (* constant 3790 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t31 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt r y0), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r y0)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt r y0) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_andi (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r y0) (l_e_st_eq_landau_n_rt_rp_5p205_t29 s0 t0 p0 p1a p1b p2 x0 i r a) ly))))))))))))). Time Defined. (* constant 3791 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t32 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt r y0), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 y0))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt r y0) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_somei l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y y0)) r (l_e_st_eq_landau_n_rt_rp_5p205_t31 s0 t0 p0 p1a p1b p2 x0 i r a y0 ly l)))))))))))))). Time Defined. (* constant 3792 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t33 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt r y0), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt r y0) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x) y0 (l_e_st_eq_landau_n_rt_rp_5p205_t32 s0 t0 p0 p1a p1b p2 x0 i r a y0 ly l)))))))))))))). Time Defined. (* constant 3793 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t34 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt r y0), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_more y0 x0))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt r y0) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_e_st_eq_landau_n_rt_satz83 x0 y0 l))))))))))))). Time Defined. (* constant 3794 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t35 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt r y0), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_and (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_more y0 x0)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt r y0) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_andi (l_e_st_eq_landau_n_rt_in y0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_more y0 x0) (l_e_st_eq_landau_n_rt_rp_5p205_t33 s0 t0 p0 p1a p1b p2 x0 i r a y0 ly l) (l_e_st_eq_landau_n_rt_rp_5p205_t34 s0 t0 p0 p1a p1b p2 x0 i r a y0 ly l)))))))))))))). Time Defined. (* constant 3795 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t36 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (ly:l_e_st_eq_landau_n_rt_lrt r y0), (forall (l:l_e_st_eq_landau_n_rt_less x0 y0), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_more y x0))))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (ly:l_e_st_eq_landau_n_rt_lrt r y0) => (fun (l:l_e_st_eq_landau_n_rt_less x0 y0) => l_somei l_e_st_eq_landau_n_rt_rat (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_more y x0)) y0 (l_e_st_eq_landau_n_rt_rp_5p205_t35 s0 t0 p0 p1a p1b p2 x0 i r a y0 ly l)))))))))))))). Time Defined. (* constant 3796 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t37 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_more y x0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0)) => l_e_st_eq_landau_n_rt_cutapp3 r x0 (l_e_st_eq_landau_n_rt_rp_5p205_t30 s0 t0 p0 p1a p1b p2 x0 i r a) (l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_more y x0))) (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt r y) => (fun (u:l_e_st_eq_landau_n_rt_less x0 y) => l_e_st_eq_landau_n_rt_rp_5p205_t36 s0 t0 p0 p1a p1b p2 x0 i r a y t u))))))))))))). Time Defined. (* constant 3797 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t38 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)), l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_more y x0)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (i:l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => l_someapp l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) (l_e_st_eq_landau_n_rt_rp_5p205_t21 s0 t0 p0 p1a p1b p2 x0 i) (l_e_st_eq_landau_n_rt_some (fun (y:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_in y (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_more y x0))) (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) => l_e_st_eq_landau_n_rt_rp_5p205_t37 s0 t0 p0 p1a p1b p2 x0 i y t)))))))))). Time Defined. (* constant 3798 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t39 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (j:l_e_st_eq_landau_n_rt_rp_in s t0), (forall (y0:l_e_st_eq_landau_n_rt_rat), (forall (uy:l_e_st_eq_landau_n_rt_urt s y0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2))))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (j:l_e_st_eq_landau_n_rt_rp_in s t0) => (fun (y0:l_e_st_eq_landau_n_rt_rat) => (fun (uy:l_e_st_eq_landau_n_rt_urt s y0) => l_e_st_eq_landau_n_rt_cut2 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2) x0 (l_e_st_eq_landau_n_rt_rp_5p205_t13 s0 t0 p0 p1a p1b p2 r i x0 lx) y0 (l_e_st_eq_landau_n_rt_rp_5p205_t20 s0 t0 p0 p1a p1b p2 s j y0 uy) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => (fun (y:l_e_st_eq_landau_n_rt_rat) => (fun (u:l_e_st_eq_landau_n_rt_less y x) => l_e_st_eq_landau_n_rt_rp_5p205_t28 s0 t0 p0 p1a p1b p2 x t y u)))) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_in x (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) => l_e_st_eq_landau_n_rt_rp_5p205_t38 s0 t0 p0 p1a p1b p2 x t)))))))))))))))). Time Defined. (* constant 3799 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t40 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (j:l_e_st_eq_landau_n_rt_rp_in s t0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (j:l_e_st_eq_landau_n_rt_rp_in s t0) => l_e_st_eq_landau_n_rt_cutapp1b s (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt s x) => l_e_st_eq_landau_n_rt_rp_5p205_t39 s0 t0 p0 p1a p1b p2 r i x0 lx s j x t)))))))))))))). Time Defined. (* constant 3800 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t41 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => l_e_st_nonemptyapp l_e_st_eq_landau_n_rt_cut t0 p1b (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_in y t0) => l_e_st_eq_landau_n_rt_rp_5p205_t40 s0 t0 p0 p1a p1b p2 r i x0 lx y t)))))))))))). Time Defined. (* constant 3801 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t42 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => l_e_st_eq_landau_n_rt_cutapp1a r (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt r x) => l_e_st_eq_landau_n_rt_rp_5p205_t41 s0 t0 p0 p1a p1b p2 r i x t)))))))))). Time Defined. (* constant 3802 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t43 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_e_st_nonemptyapp l_e_st_eq_landau_n_rt_cut s0 p1a (l_e_st_eq_landau_n_rt_cutprop (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)) (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_in y s0) => l_e_st_eq_landau_n_rt_rp_5p205_t42 s0 t0 p0 p1a p1b p2 y t)))))))). Time Defined. (* constant 3803 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_snt : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_eq_landau_n_rt_cut)))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_e_st_eq_landau_n_rt_rp_cutof (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2) (l_e_st_eq_landau_n_rt_rp_5p205_t43 s0 t0 p0 p1a p1b p2))))))). Time Defined. (* constant 3804 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t44 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => l_e_st_eq_landau_n_rt_rp_ini (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2) (l_e_st_eq_landau_n_rt_rp_5p205_t43 s0 t0 p0 p1a p1b p2) x0 lx))))))))))). Time Defined. (* constant 3805 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t45 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x0))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => l_e_st_estie l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x) x0 (l_e_st_eq_landau_n_rt_rp_5p205_t44 s0 t0 p0 p1a p1b p2 r l x0 ux lx)))))))))))). Time Defined. (* constant 3806 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t46 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)), l_e_st_eq_landau_n_rt_rp_in s s0))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0) a))))))))))))). Time Defined. (* constant 3807 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t47 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)), l_e_st_eq_landau_n_rt_lrt s x0))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0) a))))))))))))). Time Defined. (* constant 3808 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t48 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)), l_and (l_e_st_eq_landau_n_rt_urt r x0) (l_e_st_eq_landau_n_rt_lrt s x0)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)) => l_andi (l_e_st_eq_landau_n_rt_urt r x0) (l_e_st_eq_landau_n_rt_lrt s x0) ux (l_e_st_eq_landau_n_rt_rp_5p205_t47 s0 t0 p0 p1a p1b p2 r l x0 ux lx s a)))))))))))))). Time Defined. (* constant 3809 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t49 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)), l_e_st_eq_landau_n_rt_rp_less r s))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)) => l_somei l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_and (l_e_st_eq_landau_n_rt_urt r x) (l_e_st_eq_landau_n_rt_lrt s x)) x0 (l_e_st_eq_landau_n_rt_rp_5p205_t48 s0 t0 p0 p1a p1b p2 r l x0 ux lx s a)))))))))))))). Time Defined. (* constant 3810 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t50 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)), l_not (l_e_st_eq_landau_n_rt_rp_less s r)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)) => l_ec3e23 (l_e_st_eq_landau_n_rt_rp_is s r) (l_e_st_eq_landau_n_rt_rp_more s r) (l_e_st_eq_landau_n_rt_rp_less s r) (l_e_st_eq_landau_n_rt_rp_satz123b s r) (l_e_st_eq_landau_n_rt_rp_satz122 r s (l_e_st_eq_landau_n_rt_rp_5p205_t49 s0 t0 p0 p1a p1b p2 r l x0 ux lx s a))))))))))))))). Time Defined. (* constant 3811 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t51 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)), l_not (l_e_st_eq_landau_n_rt_rp_in r t0)))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_in r t0) (l_e_st_eq_landau_n_rt_rp_less s r) (l_e_st_eq_landau_n_rt_rp_5p205_t50 s0 t0 p0 p1a p1b p2 r l x0 ux lx s a) (p2 s (l_e_st_eq_landau_n_rt_rp_5p205_t46 s0 t0 p0 p1a p1b p2 r l x0 ux lx s a) r)))))))))))))). Time Defined. (* constant 3812 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t52 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)), l_e_st_eq_landau_n_rt_rp_in r s0))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_in s s0) (l_e_st_eq_landau_n_rt_lrt s x0)) => l_ore1 (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_rp_in r t0) (p0 r) (l_e_st_eq_landau_n_rt_rp_5p205_t51 s0 t0 p0 p1a p1b p2 r l x0 ux lx s a)))))))))))))). Time Defined. (* constant 3813 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t53 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (ux:l_e_st_eq_landau_n_rt_urt r x0), (forall (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), l_e_st_eq_landau_n_rt_rp_in r s0))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (ux:l_e_st_eq_landau_n_rt_urt r x0) => (fun (lx:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => l_someapp l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) (l_e_st_eq_landau_n_rt_rp_5p205_t45 s0 t0 p0 p1a p1b p2 r l x0 ux lx) (l_e_st_eq_landau_n_rt_rp_in r s0) (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) => l_e_st_eq_landau_n_rt_rp_5p205_t52 s0 t0 p0 p1a p1b p2 r l x0 ux lx y t))))))))))))). Time Defined. (* constant 3814 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t54 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), l_e_st_eq_landau_n_rt_rp_in r s0)))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => l_e_st_eq_landau_n_rt_rp_lessapp r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) l (l_e_st_eq_landau_n_rt_rp_in r s0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_urt r x) => (fun (u:l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x) => l_e_st_eq_landau_n_rt_rp_5p205_t53 s0 t0 p0 p1a p1b p2 r l x t u))))))))))). Time Defined. (* constant 3815 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t55 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), l_and (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => l_andi (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt r x0) i lx)))))))))))). Time Defined. (* constant 3816 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t56 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => l_somei l_e_st_eq_landau_n_rt_cut (fun (y:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_in y s0) (l_e_st_eq_landau_n_rt_lrt y x0)) r (l_e_st_eq_landau_n_rt_rp_5p205_t55 s0 t0 p0 p1a p1b p2 r m x0 lx ux i))))))))))))). Time Defined. (* constant 3817 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t57 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), l_e_st_eq_landau_n_rt_in x0 (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2))))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => l_e_st_estii l_e_st_eq_landau_n_rt_rat (fun (x:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_rp_5p205_schnittprop s0 t0 p0 p1a p1b p2 x) x0 (l_e_st_eq_landau_n_rt_rp_5p205_t56 s0 t0 p0 p1a p1b p2 r m x0 lx ux i))))))))))))). Time Defined. (* constant 3818 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t58 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), (forall (i:l_e_st_eq_landau_n_rt_rp_in r s0), l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r s0) => l_e_st_eq_landau_n_rt_rp_ine (l_e_st_eq_landau_n_rt_rp_5p205_schnittset s0 t0 p0 p1a p1b p2) (l_e_st_eq_landau_n_rt_rp_5p205_t43 s0 t0 p0 p1a p1b p2) x0 (l_e_st_eq_landau_n_rt_rp_5p205_t57 s0 t0 p0 p1a p1b p2 r m x0 lx ux i))))))))))))). Time Defined. (* constant 3819 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t59 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), l_not (l_e_st_eq_landau_n_rt_rp_in r s0)))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_lrt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) ux (fun (t:l_e_st_eq_landau_n_rt_rp_in r s0) => l_e_st_eq_landau_n_rt_rp_5p205_t58 s0 t0 p0 p1a p1b p2 r m x0 lx ux t)))))))))))). Time Defined. (* constant 3820 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t60 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), (forall (x0:l_e_st_eq_landau_n_rt_rat), (forall (lx:l_e_st_eq_landau_n_rt_lrt r x0), (forall (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0), l_e_st_eq_landau_n_rt_rp_in r t0))))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => (fun (x0:l_e_st_eq_landau_n_rt_rat) => (fun (lx:l_e_st_eq_landau_n_rt_lrt r x0) => (fun (ux:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x0) => l_ore2 (l_e_st_eq_landau_n_rt_rp_in r s0) (l_e_st_eq_landau_n_rt_rp_in r t0) (p0 r) (l_e_st_eq_landau_n_rt_rp_5p205_t59 s0 t0 p0 p1a p1b p2 r m x0 lx ux)))))))))))). Time Defined. (* constant 3821 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t61 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)), l_e_st_eq_landau_n_rt_rp_in r t0)))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => l_e_st_eq_landau_n_rt_rp_moreapp r (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) m (l_e_st_eq_landau_n_rt_rp_in r t0) (fun (x:l_e_st_eq_landau_n_rt_rat) => (fun (t:l_e_st_eq_landau_n_rt_lrt r x) => (fun (u:l_e_st_eq_landau_n_rt_urt (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) x) => l_e_st_eq_landau_n_rt_rp_5p205_t60 s0 t0 p0 p1a p1b p2 r m x t u))))))))))). Time Defined. (* constant 3822 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t62 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_andi (l_e_st_eq_landau_n_rt_rp_5p205_prop1 s0 t0 (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_rp_5p205_prop2 s0 t0 (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) (fun (x:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_less x (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => l_e_st_eq_landau_n_rt_rp_5p205_t54 s0 t0 p0 p1a p1b p2 x t)) (fun (x:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_more x (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2)) => l_e_st_eq_landau_n_rt_rp_5p205_t61 s0 t0 p0 p1a p1b p2 x t)))))))). Time Defined. (* constant 3823 *) Definition l_e_st_eq_landau_n_rt_rp_5p205_t63 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_eq_landau_n_rt_rp_some (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 x))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_somei l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 x) (l_e_st_eq_landau_n_rt_rp_5p205_snt s0 t0 p0 p1a p1b p2) (l_e_st_eq_landau_n_rt_rp_5p205_t62 s0 t0 p0 p1a p1b p2))))))). Time Defined. (* constant 3824 *) Definition l_e_st_eq_landau_n_rt_rp_satzp205 : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_eq_landau_n_rt_rp_one (fun (x:l_e_st_eq_landau_n_rt_cut) => l_and (l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_less y x), l_e_st_eq_landau_n_rt_rp_in y s0))) (l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_more y x), l_e_st_eq_landau_n_rt_rp_in y t0)))))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_e_onei l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 x) (l_e_st_eq_landau_n_rt_rp_5p205_t10 s0 t0 p0 p1a p1b p2) (l_e_st_eq_landau_n_rt_rp_5p205_t63 s0 t0 p0 p1a p1b p2))))))). Time Defined. (* constant 3825 *) Definition l_e_st_eq_landau_n_rt_rp_schnitt : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_eq_landau_n_rt_cut)))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_e_ind l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 x) (l_e_st_eq_landau_n_rt_rp_satzp205 s0 t0 p0 p1a p1b p2))))))). Time Defined. (* constant 3826 *) Definition l_e_st_eq_landau_n_rt_rp_satzp205a : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_less x (l_e_st_eq_landau_n_rt_rp_schnitt s0 t0 p0 p1a p1b p2)), l_e_st_eq_landau_n_rt_rp_in x s0)))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_ande1 (l_e_st_eq_landau_n_rt_rp_5p205_prop1 s0 t0 (l_e_st_eq_landau_n_rt_rp_schnitt s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_rp_5p205_prop2 s0 t0 (l_e_st_eq_landau_n_rt_rp_schnitt s0 t0 p0 p1a p1b p2)) (l_e_oneax l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 x) (l_e_st_eq_landau_n_rt_rp_satzp205 s0 t0 p0 p1a p1b p2)))))))). Time Defined. (* constant 3827 *) Definition l_e_st_eq_landau_n_rt_rp_satzp205b : (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut), (forall (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0), (forall (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))), l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_more x (l_e_st_eq_landau_n_rt_rp_schnitt s0 t0 p0 p1a p1b p2)), l_e_st_eq_landau_n_rt_rp_in x t0)))))))). exact (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (t0:l_e_st_set l_e_st_eq_landau_n_rt_cut) => (fun (p0:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => l_or (l_e_st_eq_landau_n_rt_rp_in x s0) (l_e_st_eq_landau_n_rt_rp_in x t0))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut s0) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_cut t0) => (fun (p2:l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_in x s0), l_e_st_eq_landau_n_rt_rp_all (fun (y:l_e_st_eq_landau_n_rt_cut) => (forall (u:l_e_st_eq_landau_n_rt_rp_in y t0), l_e_st_eq_landau_n_rt_rp_less x y))))) => l_ande2 (l_e_st_eq_landau_n_rt_rp_5p205_prop1 s0 t0 (l_e_st_eq_landau_n_rt_rp_schnitt s0 t0 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_rp_5p205_prop2 s0 t0 (l_e_st_eq_landau_n_rt_rp_schnitt s0 t0 p0 p1a p1b p2)) (l_e_oneax l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_5p205_prop3 s0 t0 x) (l_e_st_eq_landau_n_rt_rp_satzp205 s0 t0 p0 p1a p1b p2)))))))). Time Defined. (* constant 3828 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_i : l_e_st_eq_landau_n_rt_cut. exact l_e_st_eq_landau_n_rt_rp_1rp. Time Defined. (* constant 3829 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_r1 : (forall (r:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_ivad_i). Time Defined. (* constant 3830 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_s1 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_ivad_i)). Time Defined. (* constant 3831 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_rps : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pl r s)). Time Defined. (* constant 3832 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_2 : l_e_st_eq_landau_n_rt_cut. exact (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i). Time Defined. (* constant 3833 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t1 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) l_e_st_eq_landau_n_rt_rp_ivad_2))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pdeq12a (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s) l_e_st_eq_landau_n_rt_rp_ivad_i)). Time Defined. (* constant 3834 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t2 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_4pl23 r l_e_st_eq_landau_n_rt_rp_ivad_i s l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_asspl2 (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i))). Time Defined. (* constant 3835 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t3 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_t2 r s))). Time Defined. (* constant 3836 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t4 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_2))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_2) (l_e_st_eq_landau_n_rt_rp_ivad_t3 r s) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i))). Time Defined. (* constant 3837 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t5 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) l_e_st_eq_landau_n_rt_rp_ivad_2) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ivad_rps r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) l_e_st_eq_landau_n_rt_rp_ivad_2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_t4 r s))). Time Defined. (* constant 3838 *) Definition l_e_st_eq_landau_n_rt_rp_lemmaivad1 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl r s)) (l_e_st_eq_landau_n_rt_rp_ivad_t1 r s) (l_e_st_eq_landau_n_rt_rp_ivad_t5 r s))). Time Defined. (* constant 3839 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_rs : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_cut)). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ts r s)). Time Defined. (* constant 3840 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t6 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_tdeq12a (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s) l_e_st_eq_landau_n_rt_rp_ivad_i)). Time Defined. (* constant 3841 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t7 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) s))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) s) (l_e_st_eq_landau_n_rt_rp_disttp1 r l_e_st_eq_landau_n_rt_rp_ivad_i s) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i s) s (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) (l_e_st_eq_landau_n_rt_rp_satz151b s)))). Time Defined. (* constant 3842 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t8 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ivad_rps r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) s) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) s) (l_e_st_eq_landau_n_rt_rp_ivad_r1 r)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ivad_rps r s)) (l_e_st_eq_landau_n_rt_rp_disttp2 (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) s l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) s) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) s) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_t7 r s) (l_e_st_eq_landau_n_rt_rp_satz151 (l_e_st_eq_landau_n_rt_rp_ivad_r1 r))) (l_e_st_eq_landau_n_rt_rp_4pl24 (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) s r l_e_st_eq_landau_n_rt_rp_ivad_i))). Time Defined. (* constant 3843 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t9 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ivad_rps r s)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ivad_rps r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_t8 r s) (l_e_st_eq_landau_n_rt_rp_satz151 l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i))). Time Defined. (* constant 3844 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t10 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_ispl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i)) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_t9 r s)) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i))). Time Defined. (* constant 3845 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t11 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tr3is l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_2) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s))) (l_e_st_eq_landau_n_rt_rp_asspl1 (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_4pl23 r s l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ispl12 (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_satz151a (l_e_st_eq_landau_n_rt_rp_ivad_r1 r)) (l_e_st_eq_landau_n_rt_rp_satz151c (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s))))). Time Defined. (* constant 3846 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t12 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)))) (l_e_st_eq_landau_n_rt_rp_ivad_t10 r s) (l_e_st_eq_landau_n_rt_rp_ispl2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rps r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ivad_t11 r s)))). Time Defined. (* constant 3847 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t13 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ivad_rs r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_eqi12 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i l_e_st_eq_landau_n_rt_rp_ivad_i)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_rs r s) l_e_st_eq_landau_n_rt_rp_ivad_i) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_t12 r s))). Time Defined. (* constant 3848 *) Definition l_e_st_eq_landau_n_rt_rp_lemmaivad2 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) (l_e_st_eq_landau_n_rt_rp_df (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_pl r l_e_st_eq_landau_n_rt_rp_1rp) l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_ts l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_pl s l_e_st_eq_landau_n_rt_rp_1rp)))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts r s)) (l_e_st_eq_landau_n_rt_rp_ivad_t6 r s) (l_e_st_eq_landau_n_rt_rp_ivad_t13 r s))). Time Defined. (* constant 3849 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t14 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s) l_e_st_eq_landau_n_rt_rp_ivad_i)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) => l_e_st_eq_landau_n_rt_rp_morede12 (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s) l_e_st_eq_landau_n_rt_rp_ivad_i m))). Time Defined. (* constant 3850 *) Definition l_e_st_eq_landau_n_rt_rp_ivad_t15 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) => l_e_st_eq_landau_n_rt_rp_satz136a (l_e_st_eq_landau_n_rt_rp_ivad_r1 r) (l_e_st_eq_landau_n_rt_rp_ivad_s1 r s) l_e_st_eq_landau_n_rt_rp_ivad_i (l_e_st_eq_landau_n_rt_rp_ivad_t14 r s m)))). Time Defined. (* constant 3851 *) Definition l_e_st_eq_landau_n_rt_rp_lemmaivad3 : (forall (r:l_e_st_eq_landau_n_rt_cut), (forall (s:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)), l_e_st_eq_landau_n_rt_rp_more r s))). exact (fun (r:l_e_st_eq_landau_n_rt_cut) => (fun (s:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp r) (l_e_st_eq_landau_n_rt_rp_pdofrp s)) => l_e_st_eq_landau_n_rt_rp_satz136a r s l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_ivad_t15 r s m)))). Time Defined. (* constant 3852 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t1 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b)))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => l_e_st_eq_landau_n_rt_rp_treq2 (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b) c e f))))))). Time Defined. (* constant 3853 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t2 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td b a)) (l_e_st_eq_landau_n_rt_rp_td a a)))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td b a)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_eqpd2 (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_td a b) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_comtd b a)) (l_e_st_eq_landau_n_rt_rp_pdmd (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b))))))))). Time Defined. (* constant 3854 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t3 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b))))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => l_e_st_eq_landau_n_rt_rp_tr4eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_td b (l_e_st_eq_landau_n_rt_rp_md a b))) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_td b b))) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td b a)) (l_e_st_eq_landau_n_rt_rp_td b b)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b)) (l_e_st_eq_landau_n_rt_rp_disttpd1 a b (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_eqpd12 (l_e_st_eq_landau_n_rt_rp_td a (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td b (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_td b b)) (l_e_st_eq_landau_n_rt_rp_disttmd2 a a b) (l_e_st_eq_landau_n_rt_rp_disttmd2 b a b)) (l_e_st_eq_landau_n_rt_rp_asspd2 (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td b a) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td b b))) (l_e_st_eq_landau_n_rt_rp_eqmd1 (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td b a)) (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b) (l_e_st_eq_landau_n_rt_rp_d161_t2 c a b n o e f))))))))). Time Defined. (* constant 3855 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t4 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_md a b))))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b)) (l_e_st_eq_landau_n_rt_rp_d161_t3 c a b n o e f)) (l_e_st_eq_landau_n_rt_rp_satzd182e (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b) (l_e_st_eq_landau_n_rt_rp_d161_t1 c a b n o e f))))))))). Time Defined. (* constant 3856 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t5 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), l_or (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b))))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => l_e_st_eq_landau_n_rt_rp_satzd192c (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_md a b) (l_e_st_eq_landau_n_rt_rp_d161_t4 c a b n o e f)))))))). Time Defined. (* constant 3857 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t6 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td b b))))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_td a a) (l_e_st_eq_landau_n_rt_rp_td b b) (l_e_st_eq_landau_n_rt_rp_d161_t1 c a b n o e f) (l_e_st_eq_landau_n_rt_rp_td01 a a z))))))))). Time Defined. (* constant 3858 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t7 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_zero b)))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_zero b) (l_refimp (l_e_st_eq_landau_n_rt_rp_zero b)) (l_e_st_eq_landau_n_rt_rp_satzd192c b b (l_e_st_eq_landau_n_rt_rp_d161_t6 c a b n o e f z)))))))))). Time Defined. (* constant 3859 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t8 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_eq a b)))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_zeroeq a b z (l_e_st_eq_landau_n_rt_rp_d161_t7 c a b n o e f z))))))))). Time Defined. (* constant 3860 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t9 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_posd a)))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_negd a) (l_e_st_eq_landau_n_rt_rp_axrdo a) n p)))))))). Time Defined. (* constant 3861 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t10 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_not (l_e_st_eq_landau_n_rt_rp_zero b))))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_zero a) p (fun (t:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_d161_t7 c b a o n f e t))))))))). Time Defined. (* constant 3862 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t11 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_posd b)))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_d161_t9 c b a o n f e (l_e_st_eq_landau_n_rt_rp_d161_t10 c a b n o e f p))))))))). Time Defined. (* constant 3863 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t12 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pd a b)))))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_pnot0d (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_ppd a b (l_e_st_eq_landau_n_rt_rp_d161_t9 c a b n o e f p) (l_e_st_eq_landau_n_rt_rp_d161_t11 c a b n o e f p)))))))))). Time Defined. (* constant 3864 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t13 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b))))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_ore2 (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a b)) (l_e_st_eq_landau_n_rt_rp_d161_t5 c a b n o e f) (l_e_st_eq_landau_n_rt_rp_d161_t12 c a b n o e f p))))))))). Time Defined. (* constant 3865 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t14 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), (forall (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)), l_e_st_eq_landau_n_rt_rp_eq a b)))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => (fun (p:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_satzd182b a b (l_e_st_eq_landau_n_rt_rp_d161_t13 c a b n o e f p))))))))). Time Defined. (* constant 3866 *) Definition l_e_st_eq_landau_n_rt_rp_satzd161b : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c), (forall (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c), l_e_st_eq_landau_n_rt_rp_eq a b))))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd a)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_negd b)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a a) c) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b b) c) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_eq a b) (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_d161_t8 c a b n o e f t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero a)) => l_e_st_eq_landau_n_rt_rp_d161_t14 c a b n o e f t)))))))). Time Defined. (* constant 3867 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t15 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (z:l_e_st_eq_landau_n_rt_rp_zero c), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td c c) c))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero c) => l_e_st_eq_landau_n_rt_rp_zeroeq (l_e_st_eq_landau_n_rt_rp_td c c) c (l_e_st_eq_landau_n_rt_rp_td01 c c z) z))). Time Defined. (* constant 3868 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t16 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (z:l_e_st_eq_landau_n_rt_rp_zero c), l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd c)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td c c) c)))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero c) => l_andi (l_not (l_e_st_eq_landau_n_rt_rp_negd c)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td c c) c) n (l_e_st_eq_landau_n_rt_rp_d161_t15 c n z)))). Time Defined. (* constant 3869 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t17 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (z:l_e_st_eq_landau_n_rt_rp_zero c), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd x)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td x x) c))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero c) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd x)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td x x) c)) c (l_e_st_eq_landau_n_rt_rp_d161_t16 c n z)))). Time Defined. (* constant 3870 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t18 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)), l_e_st_eq_landau_n_rt_rp_posd c))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_zero c) (l_e_st_eq_landau_n_rt_rp_posd c) (l_e_st_eq_landau_n_rt_rp_negd c) (l_e_st_eq_landau_n_rt_rp_axrdo c) n o))). Time Defined. (* constant 3871 *) Definition l_e_st_eq_landau_n_rt_rp_d161_crp : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)), l_e_st_eq_landau_n_rt_cut))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)) => l_e_st_eq_landau_n_rt_rp_rpofpd c (l_e_st_eq_landau_n_rt_rp_d161_t18 c n o)))). Time Defined. (* constant 3872 *) Definition l_e_st_eq_landau_n_rt_rp_d161_srp : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)), l_e_st_eq_landau_n_rt_cut))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)) => l_e_st_eq_landau_n_rt_rp_sqrt (l_e_st_eq_landau_n_rt_rp_d161_crp c n o)))). Time Defined. (* constant 3873 *) Definition l_e_st_eq_landau_n_rt_rp_d161_s : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)), l_e_st_eq_landau_n_rt_rp_dif))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)) => l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_d161_srp c n o)))). Time Defined. (* constant 3874 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t19 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_d161_s c n o) (l_e_st_eq_landau_n_rt_rp_d161_s c n o)) c))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_d161_s c n o) (l_e_st_eq_landau_n_rt_rp_d161_s c n o)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_d161_srp c n o) (l_e_st_eq_landau_n_rt_rp_d161_srp c n o))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_d161_crp c n o)) c (l_e_st_eq_landau_n_rt_rp_lemmaivad2 (l_e_st_eq_landau_n_rt_rp_d161_srp c n o) (l_e_st_eq_landau_n_rt_rp_d161_srp c n o)) (l_e_st_eq_landau_n_rt_rp_isrpepd (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_d161_srp c n o) (l_e_st_eq_landau_n_rt_rp_d161_srp c n o)) (l_e_st_eq_landau_n_rt_rp_d161_crp c n o) (l_e_st_eq_landau_n_rt_rp_thsqrt1 (l_e_st_eq_landau_n_rt_rp_d161_crp c n o))) (l_e_st_eq_landau_n_rt_rp_eqpdrp2 c (l_e_st_eq_landau_n_rt_rp_d161_t18 c n o))))). Time Defined. (* constant 3875 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t20 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)), l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_d161_s c n o))) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_d161_s c n o) (l_e_st_eq_landau_n_rt_rp_d161_s c n o)) c)))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)) => l_andi (l_not (l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_d161_s c n o))) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_d161_s c n o) (l_e_st_eq_landau_n_rt_rp_d161_s c n o)) c) (l_e_st_eq_landau_n_rt_rp_pnotnd (l_e_st_eq_landau_n_rt_rp_d161_s c n o) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_d161_srp c n o))) (l_e_st_eq_landau_n_rt_rp_d161_t19 c n o)))). Time Defined. (* constant 3876 *) Definition l_e_st_eq_landau_n_rt_rp_d161_t21 : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd x)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td x x) c))))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_zero c)) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd x)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td x x) c)) (l_e_st_eq_landau_n_rt_rp_d161_s c n o) (l_e_st_eq_landau_n_rt_rp_d161_t20 c n o)))). Time Defined. (* constant 3877 *) Definition l_e_st_eq_landau_n_rt_rp_satzd161a : (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd x)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td x x) c)))). exact (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_negd c)) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_zero c) (l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd x)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td x x) c))) (fun (t:l_e_st_eq_landau_n_rt_rp_zero c) => l_e_st_eq_landau_n_rt_rp_d161_t17 c n t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero c)) => l_e_st_eq_landau_n_rt_rp_d161_t21 c n t))). Time Defined. (* constant 3878 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_absd a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_ori1 (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_absd a))) (l_e_st_eq_landau_n_rt_rp_satzd166f a z)))). Time Defined. (* constant 3879 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (n:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_absd a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) => l_e_st_eq_landau_n_rt_rp_natintd (l_e_st_eq_landau_n_rt_rp_absd a) n))). Time Defined. (* constant 3880 *) Definition l_e_st_eq_landau_n_rt_rp_intabsd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_absd a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => l_orapp (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_absd a)) i (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_intd_t1 a i t) (fun (t:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) => l_e_st_eq_landau_n_rt_rp_intd_t2 a i t))). Time Defined. (* constant 3881 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (n:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) => l_e_st_eq_landau_n_rt_rp_eqnatd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_satzd178a a) n))). Time Defined. (* constant 3882 *) Definition l_e_st_eq_landau_n_rt_rp_intm0d : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_m0d a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a))) i (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_satzd176b a t) (fun (t:l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd a)) => l_e_st_eq_landau_n_rt_rp_intd_t4 a i t))). Time Defined. (* constant 3883 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_eq b (l_e_st_eq_landau_n_rt_rp_pd a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_pd a b) b (l_e_st_eq_landau_n_rt_rp_pd01 a b z)))))). Time Defined. (* constant 3884 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_eqintd b (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_intd_t5 a b i j z) j))))). Time Defined. (* constant 3885 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t7 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_pd a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_pd a b) a (l_e_st_eq_landau_n_rt_rp_pd02 a b z)))))). Time Defined. (* constant 3886 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t8 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (z:l_e_st_eq_landau_n_rt_rp_zero b), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_eqintd a (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_intd_t7 a b i j z) i))))). Time Defined. (* constant 3887 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t9 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd a p)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_ande2 (l_e_st_eq_landau_n_rt_rp_posd a) (forall (t:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_rpofpd a t)) (l_e_st_eq_landau_n_rt_rp_posintnatd a p i) p))). Time Defined. (* constant 3888 *) Definition l_e_st_eq_landau_n_rt_rp_intd_apb1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_cut))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pd a b) pp))))). Time Defined. (* constant 3889 *) Definition l_e_st_eq_landau_n_rt_rp_intd_a1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_cut)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_rpofpd a p)))))). Time Defined. (* constant 3890 *) Definition l_e_st_eq_landau_n_rt_rp_intd_b1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_cut))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_rpofpd b q))))))). Time Defined. (* constant 3891 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t10 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_natpl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_t9 a i p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q) (l_e_st_eq_landau_n_rt_rp_intd_t9 b j q)))))))). Time Defined. (* constant 3892 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t11 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q)))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_eqpd12 a (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p)) b (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q)) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 a p) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 b q)))))))). Time Defined. (* constant 3893 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t12 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q)))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_treq (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q))) (l_e_st_eq_landau_n_rt_rp_intd_t11 a b i j pp p q) (l_e_st_eq_landau_n_rt_rp_lemmaivad1 (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q))))))))). Time Defined. (* constant 3894 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t13 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q))) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q)))) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q)) (l_e_st_eq_landau_n_rt_rp_eqpderp (l_e_st_eq_landau_n_rt_rp_pd a b) pp (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q))) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q))) (l_e_st_eq_landau_n_rt_rp_intd_t12 a b i j pp p q)) (l_e_st_eq_landau_n_rt_rp_isrppd2 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q)))))))))). Time Defined. (* constant 3895 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t14 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_isp1 l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b1 a b i j pp p q)) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp) (l_e_st_eq_landau_n_rt_rp_intd_t10 a b i j pp p q) (l_e_st_eq_landau_n_rt_rp_intd_t13 a b i j pp p q)))))))). Time Defined. (* constant 3896 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t15 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_pd a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_andi (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) (forall (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j t)) pp (fun (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_intd_t14 a b i j t p q)))))))). Time Defined. (* constant 3897 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t16 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (q:l_e_st_eq_landau_n_rt_rp_posd b), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (q:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_natintd (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_intd_t15 a b i j pp p q)))))))). Time Defined. (* constant 3898 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t17 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_m0d b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_satzd176c b n))))))). Time Defined. (* constant 3899 *) Definition l_e_st_eq_landau_n_rt_rp_intd_b2 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_cut))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_intd_t17 a b i j pp p n)))))))). Time Defined. (* constant 3900 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t18 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_md a (l_e_st_eq_landau_n_rt_rp_m0d b))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_eqpd2 b (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d b)) a (l_e_st_eq_landau_n_rt_rp_satzd177a b)))))))). Time Defined. (* constant 3901 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t19 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a (l_e_st_eq_landau_n_rt_rp_m0d b))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_md a (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_intd_t18 a b i j pp p n) pp))))))). Time Defined. (* constant 3902 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t20 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_mored a (l_e_st_eq_landau_n_rt_rp_m0d b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_satzd182a a (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_intd_t19 a b i j pp p n)))))))). Time Defined. (* constant 3903 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t21 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_eqmored12 a (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p)) (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n)) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 a p) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_intd_t17 a b i j pp p n)) (l_e_st_eq_landau_n_rt_rp_intd_t20 a b i j pp p n)))))))). Time Defined. (* constant 3904 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t22 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_lemmaivad3 (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_t21 a b i j pp p n)))))))). Time Defined. (* constant 3905 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t23 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_t22 a b i j pp p n))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_natmn (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_t9 a i p) (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_t9 (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_intm0d b j) (l_e_st_eq_landau_n_rt_rp_intd_t17 a b i j pp p n)) (l_e_st_eq_landau_n_rt_rp_intd_t22 a b i j pp p n)))))))). Time Defined. (* constant 3906 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t24 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp)))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_eqpd12 (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n)) (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp)) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_intd_t17 a b i j pp p n)) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 (l_e_st_eq_landau_n_rt_rp_pd a b) pp)))))))). Time Defined. (* constant 3907 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t25 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_eq a (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp)))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_tr4eq a (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a b) b) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp))) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a b) b) a (l_e_st_eq_landau_n_rt_rp_mdpd a b)) (l_e_st_eq_landau_n_rt_rp_compd (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_intd_t24 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_lemmaivad1 (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp))))))))). Time Defined. (* constant 3908 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t26 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp)) (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp)) (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp))) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp)))) (l_e_st_eq_landau_n_rt_rp_isrppd1 (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp))) (l_e_st_eq_landau_n_rt_rp_eqpderp a p (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp))) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp))) (l_e_st_eq_landau_n_rt_rp_intd_t25 a b i j pp p n))))))))). Time Defined. (* constant 3909 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t27 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_t22 a b i j pp p n))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_satz140g (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp) (l_e_st_eq_landau_n_rt_rp_intd_t22 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_t26 a b i j pp p n)))))))). Time Defined. (* constant 3910 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t28 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_isp1 l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) (l_e_st_eq_landau_n_rt_rp_mn (l_e_st_eq_landau_n_rt_rp_intd_a1 a b i j pp p) (l_e_st_eq_landau_n_rt_rp_intd_b2 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_t22 a b i j pp p n)) (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j pp) (l_e_st_eq_landau_n_rt_rp_intd_t23 a b i j pp p n) (l_e_st_eq_landau_n_rt_rp_intd_t27 a b i j pp p n)))))))). Time Defined. (* constant 3911 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t29 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_pd a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_andi (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) (forall (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_intd_apb1 a b i j t)) pp (fun (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_intd_t28 a b i j t p n)))))))). Time Defined. (* constant 3912 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t30 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), (forall (n:l_e_st_eq_landau_n_rt_rp_negd b), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_natintd (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_intd_t29 a b i j pp p n)))))))). Time Defined. (* constant 3913 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t31 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_rappd b (l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd b) => l_e_st_eq_landau_n_rt_rp_intd_t16 a b i j pp p t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_intd_t8 a b i j t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_intd_t30 a b i j pp p t))))))). Time Defined. (* constant 3914 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t31a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_not (l_e_st_eq_landau_n_rt_rp_negd b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_pnotnd (l_e_st_eq_landau_n_rt_rp_pd a b) pp) (fun (t:l_e_st_eq_landau_n_rt_rp_negd b) => l_e_st_eq_landau_n_rt_rp_npd a b n t))))))). Time Defined. (* constant 3915 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t32 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_not (l_e_st_eq_landau_n_rt_rp_zero b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_pnotnd (l_e_st_eq_landau_n_rt_rp_pd a b) pp) (fun (t:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_eqnegd a (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_symeq (l_e_st_eq_landau_n_rt_rp_pd a b) a (l_e_st_eq_landau_n_rt_rp_pd02 a b t)) n))))))). Time Defined. (* constant 3916 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t33 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_posd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_negd b) (l_e_st_eq_landau_n_rt_rp_axrdo b) (l_e_st_eq_landau_n_rt_rp_intd_t31a a b i j pp n) (l_e_st_eq_landau_n_rt_rp_intd_t32 a b i j pp n))))))). Time Defined. (* constant 3917 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t34 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd b a))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_pd b a) (l_e_st_eq_landau_n_rt_rp_compd a b) pp)))))). Time Defined. (* constant 3918 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t35 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd b a))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_intd_t30 b a j i (l_e_st_eq_landau_n_rt_rp_intd_t34 a b i j pp n) (l_e_st_eq_landau_n_rt_rp_intd_t33 a b i j pp n) n)))))). Time Defined. (* constant 3919 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t36 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_eqintd (l_e_st_eq_landau_n_rt_rp_pd b a) (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_compd b a) (l_e_st_eq_landau_n_rt_rp_intd_t35 a b i j pp n))))))). Time Defined. (* constant 3920 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t37 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (pp:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_rappd a (l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a) => l_e_st_eq_landau_n_rt_rp_intd_t31 a b i j pp t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_intd_t6 a b i j t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a) => l_e_st_eq_landau_n_rt_rp_intd_t36 a b i j pp t)))))). Time Defined. (* constant 3921 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t38 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n0p:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n0p:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_intdi0 (l_e_st_eq_landau_n_rt_rp_pd a b) n0p))))). Time Defined. (* constant 3922 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t39 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_satzd176c (l_e_st_eq_landau_n_rt_rp_pd a b) np))))). Time Defined. (* constant 3923 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t40 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_eqposd (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_satzd180 a b) (l_e_st_eq_landau_n_rt_rp_intd_t39 a b i j np)))))). Time Defined. (* constant 3924 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t41 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_intd_t37 (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_intm0d a i) (l_e_st_eq_landau_n_rt_rp_intm0d b j) (l_e_st_eq_landau_n_rt_rp_intd_t40 a b i j np)))))). Time Defined. (* constant 3925 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t42 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_eqintd (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_satzd180a a b) (l_e_st_eq_landau_n_rt_rp_intd_t41 a b i j np)))))). Time Defined. (* constant 3926 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t43 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_intm0d (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_intd_t42 a b i j np)))))). Time Defined. (* constant 3927 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t44 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (np:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_eqintd (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a b))) (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_satzd177 (l_e_st_eq_landau_n_rt_rp_pd a b)) (l_e_st_eq_landau_n_rt_rp_intd_t43 a b i j np)))))). Time Defined. (* constant 3928 *) Definition l_e_st_eq_landau_n_rt_rp_intpd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => l_e_st_eq_landau_n_rt_rp_rappd (l_e_st_eq_landau_n_rt_rp_pd a b) (l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a b)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_intd_t37 a b i j t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_intd_t38 a b i j t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a b)) => l_e_st_eq_landau_n_rt_rp_intd_t44 a b i j t))))). Time Defined. (* constant 3929 *) Definition l_e_st_eq_landau_n_rt_rp_intmd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_md a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => l_e_st_eq_landau_n_rt_rp_intpd a (l_e_st_eq_landau_n_rt_rp_m0d b) i (l_e_st_eq_landau_n_rt_rp_intm0d b j))))). Time Defined. (* constant 3930 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t45 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_not (l_e_st_eq_landau_n_rt_rp_zero a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_zero a) (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)) n (fun (t:l_e_st_eq_landau_n_rt_rp_zero a) => l_e_st_eq_landau_n_rt_rp_td01 a b t)))))). Time Defined. (* constant 3931 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t46 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_not (l_e_st_eq_landau_n_rt_rp_zero b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_zero b) (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b)) n (fun (t:l_e_st_eq_landau_n_rt_rp_zero b) => l_e_st_eq_landau_n_rt_rp_td02 a b t)))))). Time Defined. (* constant 3932 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t47 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_satzd166e a (l_e_st_eq_landau_n_rt_rp_intd_t45 a b i j n)))))). Time Defined. (* constant 3933 *) Definition l_e_st_eq_landau_n_rt_rp_intd_a3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_cut))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_intd_t47 a b i j n)))))). Time Defined. (* constant 3934 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t48 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd b)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_satzd166e b (l_e_st_eq_landau_n_rt_rp_intd_t46 a b i j n)))))). Time Defined. (* constant 3935 *) Definition l_e_st_eq_landau_n_rt_rp_intd_b3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_cut))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_intd_t48 a b i j n)))))). Time Defined. (* constant 3936 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t49 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_natts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_t9 (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_intabsd a i) (l_e_st_eq_landau_n_rt_rp_intd_t47 a b i j n)) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_t9 (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_intabsd b j) (l_e_st_eq_landau_n_rt_rp_intd_t48 a b i j n))))))). Time Defined. (* constant 3937 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t50 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_satzd166e (l_e_st_eq_landau_n_rt_rp_td a b) n))))). Time Defined. (* constant 3938 *) Definition l_e_st_eq_landau_n_rt_rp_intd_atb3 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_cut)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) p)))))). Time Defined. (* constant 3939 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t51 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_eqtd12 (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n)) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n)) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_intd_t47 a b i j n)) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_intd_t48 a b i j n)))))))). Time Defined. (* constant 3940 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t52 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_tr3eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))) (l_e_st_eq_landau_n_rt_rp_satzd193 a b) (l_e_st_eq_landau_n_rt_rp_intd_t51 a b i j n p) (l_e_st_eq_landau_n_rt_rp_lemmaivad2 (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n)))))))). Time Defined. (* constant 3941 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t53 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_intd_atb3 a b i j n p) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n)))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_tris2 l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_intd_atb3 a b i j n p) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n)) (l_e_st_eq_landau_n_rt_rp_rpofpd (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n)))) (l_e_st_eq_landau_n_rt_rp_eqpderp (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b)) p (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))) (l_e_st_eq_landau_n_rt_rp_posdirp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))) (l_e_st_eq_landau_n_rt_rp_intd_t52 a b i j n p)) (l_e_st_eq_landau_n_rt_rp_isrppd1 (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n))))))))). Time Defined. (* constant 3942 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t54 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), (forall (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_intd_atb3 a b i j n p))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_isp1 l_e_st_eq_landau_n_rt_cut (fun (t:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_natrp t) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_intd_a3 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_b3 a b i j n)) (l_e_st_eq_landau_n_rt_rp_intd_atb3 a b i j n p) (l_e_st_eq_landau_n_rt_rp_intd_t49 a b i j n) (l_e_st_eq_landau_n_rt_rp_intd_t53 a b i j n p))))))). Time Defined. (* constant 3943 *) Definition l_e_st_eq_landau_n_rt_rp_intd_t55 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_natd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_andi (l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))) (forall (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_intd_atb3 a b i j n t)) (l_e_st_eq_landau_n_rt_rp_intd_t50 a b i j n) (fun (t:l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_intd_t54 a b i j n t)))))). Time Defined. (* constant 3944 *) Definition l_e_st_eq_landau_n_rt_rp_inttd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a), (forall (j:l_e_st_eq_landau_n_rt_rp_intd b), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_td a b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a) => (fun (j:l_e_st_eq_landau_n_rt_rp_intd b) => (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a b))) => l_e_st_eq_landau_n_rt_rp_intd_t55 a b i j t))))). Time Defined. (* constant 3945 *) Definition l_e_st_eq_landau_n_rt_rp_r_eq : (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq x y)). Time Defined. (* constant 3946 *) Definition l_e_st_eq_landau_n_rt_rp_r_refeq : (forall (x:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_r_eq x x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_refeq x). Time Defined. (* constant 3947 *) Definition l_e_st_eq_landau_n_rt_rp_r_symeq : (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (t:l_e_st_eq_landau_n_rt_rp_r_eq x y), l_e_st_eq_landau_n_rt_rp_r_eq y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_eq x y) => l_e_st_eq_landau_n_rt_rp_symeq x y t))). Time Defined. (* constant 3948 *) Definition l_e_st_eq_landau_n_rt_rp_r_treq : (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (t:l_e_st_eq_landau_n_rt_rp_r_eq x y), (forall (u:l_e_st_eq_landau_n_rt_rp_r_eq y z), l_e_st_eq_landau_n_rt_rp_r_eq x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_eq x y) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_eq y z) => l_e_st_eq_landau_n_rt_rp_treq x y z t u))))). Time Defined. (* constant 3949 *) Definition l_e_st_eq_landau_n_rt_rp_r_inn : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_esti l_e_st_eq_landau_n_rt_rp_dif a s)). Time Defined. (* constant 3950 *) Definition l_e_st_eq_landau_n_rt_rp_r_real : Type. exact (l_e_st_eq_ect l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq). Time Defined. (* constant 3951 *) Definition l_e_st_eq_landau_n_rt_rp_r_is : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_is l_e_st_eq_landau_n_rt_rp_r_real r s)). Time Defined. (* constant 3952 *) Definition l_e_st_eq_landau_n_rt_rp_r_nis : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_not (l_e_st_eq_landau_n_rt_rp_r_is r s))). Time Defined. (* constant 3953 *) Definition l_e_st_eq_landau_n_rt_rp_r_some : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_r_real), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_r_real), Prop)) => l_some l_e_st_eq_landau_n_rt_rp_r_real p). Time Defined. (* constant 3954 *) Definition l_e_st_eq_landau_n_rt_rp_r_all : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_r_real), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_r_real), Prop)) => l_all l_e_st_eq_landau_n_rt_rp_r_real p). Time Defined. (* constant 3955 *) Definition l_e_st_eq_landau_n_rt_rp_r_one : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rp_r_real), Prop)), Prop). exact (fun (p:(forall (x:l_e_st_eq_landau_n_rt_rp_r_real), Prop)) => l_e_one l_e_st_eq_landau_n_rt_rp_r_real p). Time Defined. (* constant 3956 *) Definition l_e_st_eq_landau_n_rt_rp_r_in : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s0:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s0:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_real r s0)). Time Defined. (* constant 3957 *) Definition l_e_st_eq_landau_n_rt_rp_r_realof : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_ectelt l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq a). Time Defined. (* constant 3958 *) Definition l_e_st_eq_landau_n_rt_rp_r_class : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_set l_e_st_eq_landau_n_rt_rp_dif). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_ecect l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq r). Time Defined. (* constant 3959 *) Definition l_e_st_eq_landau_n_rt_rp_r_innclass : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class (l_e_st_eq_landau_n_rt_rp_r_realof a))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_4_th5 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq a). Time Defined. (* constant 3960 *) Definition l_e_st_eq_landau_n_rt_rp_r_eqinn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_4_th8 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq r a air b e))))). Time Defined. (* constant 3961 *) Definition l_e_st_eq_landau_n_rt_rp_r_realapp1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), p))), p))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), p))) => l_e_st_eq_4_th3 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq r p p1))). Time Defined. (* constant 3962 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), p))))), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), p)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), p))))) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_realapp1 s p (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => p1 a y air))))))). Time Defined. (* constant 3963 *) Definition l_e_st_eq_landau_n_rt_rp_r_realapp2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), p))))), p)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), p))))) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r p (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t1 r s p p1 x xi)))))). Time Defined. (* constant 3964 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)), p))))))), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), p))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)), p))))))) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_realapp2 s t p (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => p1 a y z air))))))))). Time Defined. (* constant 3965 *) Definition l_e_st_eq_landau_n_rt_rp_r_realapp3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)), p))))))), p))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)), p))))))) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r p (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t2 r s t p p1 x xi))))))). Time Defined. (* constant 3966 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (v:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (vi:l_e_st_eq_landau_n_rt_rp_r_inn v (l_e_st_eq_landau_n_rt_rp_r_class u)), p))))))))), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), p)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (v:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (vi:l_e_st_eq_landau_n_rt_rp_r_inn v (l_e_st_eq_landau_n_rt_rp_r_class u)), p))))))))) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_realapp3 s t u p (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (v:l_e_st_eq_landau_n_rt_rp_dif) => p1 a y z v air))))))))))). Time Defined. (* constant 3967 *) Definition l_e_st_eq_landau_n_rt_rp_r_realapp4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:Prop), (forall (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (v:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (vi:l_e_st_eq_landau_n_rt_rp_r_inn v (l_e_st_eq_landau_n_rt_rp_r_class u)), p))))))))), p)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:Prop) => (fun (p1:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), (forall (z:l_e_st_eq_landau_n_rt_rp_dif), (forall (v:l_e_st_eq_landau_n_rt_rp_dif), (forall (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (vi:l_e_st_eq_landau_n_rt_rp_r_inn v (l_e_st_eq_landau_n_rt_rp_r_class u)), p))))))))) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r p (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t3 r s t u p p1 x xi)))))))). Time Defined. (* constant 3968 *) Definition l_e_st_eq_landau_n_rt_rp_r_isin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a1 b1), l_e_st_eq_landau_n_rt_rp_r_is r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a1 b1) => l_e_st_eq_5_th3 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq r s a1 a1ir b1 b1is e))))))). Time Defined. (* constant 3969 *) Definition l_e_st_eq_landau_n_rt_rp_r_isex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_eq a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_5_th5 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq r s a1 a1ir b1 b1is i))))))). Time Defined. (* constant 3970 *) Definition l_e_st_eq_landau_n_rt_rp_r_nisin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_eq a1 b1)), l_e_st_eq_landau_n_rt_rp_r_nis r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_eq a1 b1)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_eq a1 b1) n (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_isex r s a1 b1 a1ir b1is t)))))))). Time Defined. (* constant 3971 *) Definition l_e_st_eq_landau_n_rt_rp_r_nisex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis r s), l_not (l_e_st_eq_landau_n_rt_rp_eq a1 b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis r s) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_r_is r s) n (fun (t:l_e_st_eq_landau_n_rt_rp_eq a1 b1) => l_e_st_eq_landau_n_rt_rp_r_isin r s a1 b1 a1ir b1is t)))))))). Time Defined. (* constant 3972 *) Definition l_e_st_eq_landau_n_rt_rp_r_fixf : (forall (alpha:Type), (forall (f:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), alpha)), Prop)). exact (fun (alpha:Type) => (fun (f:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), alpha)) => l_e_st_eq_fixfu l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq alpha f)). Time Defined. (* constant 3973 *) Definition l_e_st_eq_landau_n_rt_rp_r_indreal : (forall (alpha:Type), (forall (f:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), alpha)), (forall (ff:l_e_st_eq_landau_n_rt_rp_r_fixf alpha f), (forall (r0:l_e_st_eq_landau_n_rt_rp_r_real), alpha)))). exact (fun (alpha:Type) => (fun (f:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), alpha)) => (fun (ff:l_e_st_eq_landau_n_rt_rp_r_fixf alpha f) => (fun (r0:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_indeq l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq alpha f ff r0)))). Time Defined. (* constant 3974 *) Definition l_e_st_eq_landau_n_rt_rp_r_isindreal : (forall (alpha:Type), (forall (f:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), alpha)), (forall (ff:l_e_st_eq_landau_n_rt_rp_r_fixf alpha f), (forall (r0:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r0)), l_e_is alpha (f a) (l_e_st_eq_landau_n_rt_rp_r_indreal alpha f ff r0))))))). exact (fun (alpha:Type) => (fun (f:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), alpha)) => (fun (ff:l_e_st_eq_landau_n_rt_rp_r_fixf alpha f) => (fun (r0:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r0)) => l_e_st_eq_10_th2 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq alpha f ff r0 a air)))))). Time Defined. (* constant 3975 *) Definition l_e_st_eq_landau_n_rt_rp_r_fixf2 : (forall (alpha:Type), (forall (g:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), alpha))), Prop)). exact (fun (alpha:Type) => (fun (g:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), alpha))) => l_e_st_eq_fixfu2 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq alpha g)). Time Defined. (* constant 3976 *) Definition l_e_st_eq_landau_n_rt_rp_r_indreal2 : (forall (alpha:Type), (forall (g:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), alpha))), (forall (ff2:l_e_st_eq_landau_n_rt_rp_r_fixf2 alpha g), (forall (r0:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s0:l_e_st_eq_landau_n_rt_rp_r_real), alpha))))). exact (fun (alpha:Type) => (fun (g:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), alpha))) => (fun (ff2:l_e_st_eq_landau_n_rt_rp_r_fixf2 alpha g) => (fun (r0:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s0:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_indeq2 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq alpha g ff2 r0 s0))))). Time Defined. (* constant 3977 *) Definition l_e_st_eq_landau_n_rt_rp_r_isindreal2 : (forall (alpha:Type), (forall (g:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), alpha))), (forall (ff2:l_e_st_eq_landau_n_rt_rp_r_fixf2 alpha g), (forall (r0:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s0:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r0)), (forall (bis:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s0)), l_e_is alpha (g a b) (l_e_st_eq_landau_n_rt_rp_r_indreal2 alpha g ff2 r0 s0)))))))))). exact (fun (alpha:Type) => (fun (g:(forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), alpha))) => (fun (ff2:l_e_st_eq_landau_n_rt_rp_r_fixf2 alpha g) => (fun (r0:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s0:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r0)) => (fun (bis:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s0)) => l_e_st_eq_11_th1 l_e_st_eq_landau_n_rt_rp_dif l_e_st_eq_landau_n_rt_rp_r_eq l_e_st_eq_landau_n_rt_rp_r_refeq l_e_st_eq_landau_n_rt_rp_r_symeq l_e_st_eq_landau_n_rt_rp_r_treq alpha g ff2 r0 s0 a air b bis))))))))). Time Defined. (* constant 3978 *) Definition l_e_st_eq_landau_n_rt_rp_r_0 : l_e_st_eq_landau_n_rt_rp_r_real. exact (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)). Time Defined. (* constant 3979 *) Definition l_e_st_eq_landau_n_rt_rp_r_0in : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a0), l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a0) => l_e_st_eq_landau_n_rt_rp_r_isin r l_e_st_eq_landau_n_rt_rp_r_0 a0 (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) a0ir (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_zeroeq a0 (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) z (l_e_st_eq_landau_n_rt_rp_zeroi l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp (l_e_refis l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_rp_1rp))))))). Time Defined. (* constant 3980 *) Definition l_e_st_eq_landau_n_rt_rp_r_0ex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_zero a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_eqzero (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) a0 (l_e_st_eq_landau_n_rt_rp_r_isex l_e_st_eq_landau_n_rt_rp_r_0 r (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) a0 (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) a0ir (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r l_e_st_eq_landau_n_rt_rp_r_0 i)) (l_e_tris l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_stm (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) l_e_st_eq_landau_n_rt_rp_1rp (l_e_st_eq_landau_n_rt_rp_std (l_e_st_eq_landau_n_rt_rp_df l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)) (l_e_st_eq_landau_n_rt_rp_stmis l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp) (l_e_st_eq_landau_n_rt_rp_isstd l_e_st_eq_landau_n_rt_rp_1rp l_e_st_eq_landau_n_rt_rp_1rp)))))). Time Defined. (* constant 3981 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_propp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_posd a0))). Time Defined. (* constant 3982 *) Definition l_e_st_eq_landau_n_rt_rp_r_pos : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r x)). Time Defined. (* constant 3983 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a0), l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a0) => l_andi (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_posd a0) a0ir p)))). Time Defined. (* constant 3984 *) Definition l_e_st_eq_landau_n_rt_rp_r_posin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a0), l_e_st_eq_landau_n_rt_rp_r_pos r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a0) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r x) a0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t4 r a0 a0ir p))))). Time Defined. (* constant 3985 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (q1:l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r a), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (q1:l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r a) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_posd a) q1)))))). Time Defined. (* constant 3986 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (q1:l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r a), l_e_st_eq_landau_n_rt_rp_posd a)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (q1:l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r a) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_posd a) q1)))))). Time Defined. (* constant 3987 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (q1:l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r a), l_e_st_eq_landau_n_rt_rp_posd a0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (q1:l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r a) => l_e_st_eq_landau_n_rt_rp_eqposd a a0 (l_e_st_eq_landau_n_rt_rp_r_isex r r a a0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t5 r a0 a0ir p a q1) a0ir (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t6 r a0 a0ir p a q1))))))). Time Defined. (* constant 3988 *) Definition l_e_st_eq_landau_n_rt_rp_r_posex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_posd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r x) p (l_e_st_eq_landau_n_rt_rp_posd a0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_ivr1_propp r x) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t7 r a0 a0ir p x t)))))). Time Defined. (* constant 3989 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_propn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_negd a0))). Time Defined. (* constant 3990 *) Definition l_e_st_eq_landau_n_rt_rp_r_neg : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r x)). Time Defined. (* constant 3991 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a0), l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a0) => l_andi (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_negd a0) a0ir n)))). Time Defined. (* constant 3992 *) Definition l_e_st_eq_landau_n_rt_rp_r_negin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a0), l_e_st_eq_landau_n_rt_rp_r_neg r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a0) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r x) a0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t8 r a0 a0ir n))))). Time Defined. (* constant 3993 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (pl:l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r a), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (pl:l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r a) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_negd a) pl)))))). Time Defined. (* constant 3994 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (pl:l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r a), l_e_st_eq_landau_n_rt_rp_negd a)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (pl:l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r a) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_negd a) pl)))))). Time Defined. (* constant 3995 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (pl:l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r a), l_e_st_eq_landau_n_rt_rp_negd a0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (pl:l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r a) => l_e_st_eq_landau_n_rt_rp_eqnegd a a0 (l_e_st_eq_landau_n_rt_rp_r_isex r r a a0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t9 r a0 a0ir n a pl) a0ir (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t10 r a0 a0ir n a pl))))))). Time Defined. (* constant 3996 *) Definition l_e_st_eq_landau_n_rt_rp_r_negex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_negd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r x) n (l_e_st_eq_landau_n_rt_rp_negd a0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_ivr1_propn r x) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t11 r a0 a0ir n x t)))))). Time Defined. (* constant 3997 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_posd a0), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_posd a0) => l_or3i2 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_posin r a0 a0ir p))))). Time Defined. (* constant 3998 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t13 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (z:l_e_st_eq_landau_n_rt_rp_zero a0), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (z:l_e_st_eq_landau_n_rt_rp_zero a0) => l_or3i1 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_0in r a0 a0ir z))))). Time Defined. (* constant 3999 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t14 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_negd a0), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_negd a0) => l_or3i3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_negin r a0 a0ir n))))). Time Defined. (* constant 4000 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t15 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_rappd a0 (l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r)) (fun (t:l_e_st_eq_landau_n_rt_rp_posd a0) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t12 r a0 a0ir t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero a0) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t13 r a0 a0ir t) (fun (t:l_e_st_eq_landau_n_rt_rp_negd a0) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t14 r a0 a0ir t)))). Time Defined. (* constant 4001 *) Definition l_e_st_eq_landau_n_rt_rp_r_axrlo : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t15 r x xi))). Time Defined. (* constant 4002 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t16 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_not (l_e_st_eq_landau_n_rt_rp_r_pos r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_posd a0) (l_e_st_eq_landau_n_rt_rp_0notpd a0 (l_e_st_eq_landau_n_rt_rp_r_0ex r a0 a0ir i)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_posex r a0 a0ir t))))). Time Defined. (* constant 4003 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t17 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_not (l_e_st_eq_landau_n_rt_rp_r_neg r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_negd a0) (l_e_st_eq_landau_n_rt_rp_pnotnd a0 (l_e_st_eq_landau_n_rt_rp_r_posex r a0 a0ir p)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_negex r a0 a0ir t))))). Time Defined. (* constant 4004 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t18 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_zero a0) (l_e_st_eq_landau_n_rt_rp_nnot0d a0 (l_e_st_eq_landau_n_rt_rp_r_negex r a0 a0ir n)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_0ex r a0 a0ir t))))). Time Defined. (* constant 4005 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t19 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_ec3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_ec3_th6 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t16 r a0 a0ir t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t17 r a0 a0ir t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t18 r a0 a0ir t)))). Time Defined. (* constant 4006 *) Definition l_e_st_eq_landau_n_rt_rp_r_axrle : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_ec3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_ec3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t19 r x xi))). Time Defined. (* constant 4007 *) Definition l_e_st_eq_landau_n_rt_rp_r_axrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_orec3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_orec3i (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_axrlo r) (l_e_st_eq_landau_n_rt_rp_r_axrle r)). Time Defined. (* constant 4008 *) Definition l_e_st_eq_landau_n_rt_rp_r_rapp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:Prop), (forall (p1:(forall (t:l_e_st_eq_landau_n_rt_rp_r_pos r), p)), (forall (p2:(forall (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), p)), (forall (p3:(forall (t:l_e_st_eq_landau_n_rt_rp_r_neg r), p)), p))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:Prop) => (fun (p1:(forall (t:l_e_st_eq_landau_n_rt_rp_r_pos r), p)) => (fun (p2:(forall (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), p)) => (fun (p3:(forall (t:l_e_st_eq_landau_n_rt_rp_r_neg r), p)) => l_or3app (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) p (l_e_st_eq_landau_n_rt_rp_r_axrlo r) p2 p1 p3))))). Time Defined. (* constant 4009 *) Definition l_e_st_eq_landau_n_rt_rp_r_pnotn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_not (l_e_st_eq_landau_n_rt_rp_r_neg r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_ec3e23 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_axrle r) p)). Time Defined. (* constant 4010 *) Definition l_e_st_eq_landau_n_rt_rp_r_pnot0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_ec3e21 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_axrle r) p)). Time Defined. (* constant 4011 *) Definition l_e_st_eq_landau_n_rt_rp_r_0notp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_not (l_e_st_eq_landau_n_rt_rp_r_pos r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_ec3e12 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_axrle r) i)). Time Defined. (* constant 4012 *) Definition l_e_st_eq_landau_n_rt_rp_r_0notn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_not (l_e_st_eq_landau_n_rt_rp_r_neg r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_ec3e13 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_axrle r) i)). Time Defined. (* constant 4013 *) Definition l_e_st_eq_landau_n_rt_rp_r_nnotp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_not (l_e_st_eq_landau_n_rt_rp_r_pos r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_ec3e32 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_axrle r) n)). Time Defined. (* constant 4014 *) Definition l_e_st_eq_landau_n_rt_rp_r_nnot0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_ec3e31 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_axrle r) n)). Time Defined. (* constant 4015 *) Definition l_e_st_eq_landau_n_rt_rp_r_ispos : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_pos s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pos x) r s p i)))). Time Defined. (* constant 4016 *) Definition l_e_st_eq_landau_n_rt_rp_r_isneg : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_neg s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_neg x) r s n i)))). Time Defined. (* constant 4017 *) Definition l_e_st_eq_landau_n_rt_rp_r_pofrp : (forall (r0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_pdofrp r0)). Time Defined. (* constant 4018 *) Definition l_e_st_eq_landau_n_rt_rp_r_nofrp : (forall (r0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_ndofrp r0)). Time Defined. (* constant 4019 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrpep : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is r0 s0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is r0 s0) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_pofrp x) r0 s0 i))). Time Defined. (* constant 4020 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrpen : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_is r0 s0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_nofrp s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_is r0 s0) => l_e_isf l_e_st_eq_landau_n_rt_cut l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_nofrp x) r0 s0 i))). Time Defined. (* constant 4021 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t20 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp s0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp s0)) i))). Time Defined. (* constant 4022 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrpip : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)), l_e_st_eq_landau_n_rt_rp_is r0 s0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_isrpipd r0 s0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t20 r0 s0 i)))). Time Defined. (* constant 4023 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t21 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_nofrp s0)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_ndofrp r0) (l_e_st_eq_landau_n_rt_rp_ndofrp s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_nofrp s0)) => l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_nofrp s0) (l_e_st_eq_landau_n_rt_rp_ndofrp r0) (l_e_st_eq_landau_n_rt_rp_ndofrp s0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_ndofrp r0)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_ndofrp s0)) i))). Time Defined. (* constant 4024 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrpin : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_nofrp s0)), l_e_st_eq_landau_n_rt_rp_is r0 s0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_nofrp s0)) => l_e_st_eq_landau_n_rt_rp_isrpind r0 s0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t21 r0 s0 i)))). Time Defined. (* constant 4025 *) Definition l_e_st_eq_landau_n_rt_rp_r_posi : (forall (r0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pofrp r0)). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_posdirp r0)). Time Defined. (* constant 4026 *) Definition l_e_st_eq_landau_n_rt_rp_r_negi : (forall (r0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_nofrp r0)). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_negin (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_ndofrp r0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_ndofrp r0)) (l_e_st_eq_landau_n_rt_rp_negdirp r0)). Time Defined. (* constant 4027 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t22 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp r0)), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is s (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)), (forall (k:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_is r0 s0))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp r0)) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is s (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_isrpip r0 s0 (l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) r s (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) i) k j)))))))). Time Defined. (* constant 4028 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t23 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_amone l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_cut) => (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp y)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t22 r r x y t u (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)))))). Time Defined. (* constant 4029 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t24 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_posd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_posex r a0 a0ir p1)))). Time Defined. (* constant 4030 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_pr : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_cut)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_rpofpd a0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t24 r a0 a0ir p1))))). Time Defined. (* constant 4031 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t25 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_ivr1_pr r a0 a0ir p1)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_isin r (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_ivr1_pr r a0 a0ir p1)) a0 (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_r_ivr1_pr r a0 a0ir p1)) a0ir (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_r_ivr1_pr r a0 a0ir p1))) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 a0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t24 r a0 a0ir p1)))))). Time Defined. (* constant 4032 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t26 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_some (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_somei l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_pr r a0 a0ir p1) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t25 r a0 a0ir p1))))). Time Defined. (* constant 4033 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t27 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_some (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_some (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t26 r x t p)))). Time Defined. (* constant 4034 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t28 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_one (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_onei l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t23 r) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t27 r p))). Time Defined. (* constant 4035 *) Definition l_e_st_eq_landau_n_rt_rp_r_rpofp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_cut)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_ind l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t28 r p))). Time Defined. (* constant 4036 *) Definition l_e_st_eq_landau_n_rt_rp_r_isprp1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp r p)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_oneax l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_pofrp x)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t28 r p))). Time Defined. (* constant 4037 *) Definition l_e_st_eq_landau_n_rt_rp_r_isprp2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp r p)) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp r p)) (l_e_st_eq_landau_n_rt_rp_r_isprp1 r p))). Time Defined. (* constant 4038 *) Definition l_e_st_eq_landau_n_rt_rp_r_isperp : (forall (r1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r1), (forall (s1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s1), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r1 s1), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_r_rpofp r1 p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s1 q)))))). exact (fun (r1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r1) => (fun (s1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s1) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r1 s1) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t22 r1 s1 (l_e_st_eq_landau_n_rt_rp_r_rpofp r1 p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s1 q) (l_e_st_eq_landau_n_rt_rp_r_isprp1 r1 p) (l_e_st_eq_landau_n_rt_rp_r_isprp1 s1 q) i))))). Time Defined. (* constant 4039 *) Definition l_e_st_eq_landau_n_rt_rp_r_ispirp : (forall (r1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r1), (forall (s1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s1), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_r_rpofp r1 p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s1 q)), l_e_st_eq_landau_n_rt_rp_r_is r1 s1))))). exact (fun (r1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r1) => (fun (s1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s1) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_r_rpofp r1 p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s1 q)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real r1 (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp r1 p)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp s1 q)) s1 (l_e_st_eq_landau_n_rt_rp_r_isprp1 r1 p) (l_e_st_eq_landau_n_rt_rp_r_isrpep (l_e_st_eq_landau_n_rt_rp_r_rpofp r1 p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s1 q) i) (l_e_st_eq_landau_n_rt_rp_r_isprp2 s1 q)))))). Time Defined. (* constant 4040 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrpp1 : (forall (r0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is r0 (l_e_st_eq_landau_n_rt_rp_r_rpofp (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_posi r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t22 (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) r0 (l_e_st_eq_landau_n_rt_rp_r_rpofp (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_posi r0)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pofrp r0)) (l_e_st_eq_landau_n_rt_rp_r_isprp1 (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_posi r0)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pofrp r0))). Time Defined. (* constant 4041 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrpp2 : (forall (r0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_r_rpofp (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_posi r0)) r0). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut r0 (l_e_st_eq_landau_n_rt_rp_r_rpofp (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_posi r0)) (l_e_st_eq_landau_n_rt_rp_r_isrpp1 r0)). Time Defined. (* constant 4042 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t29 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp r0)), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is s (l_e_st_eq_landau_n_rt_rp_r_nofrp s0)), (forall (k:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_is r0 s0))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp r0)) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is s (l_e_st_eq_landau_n_rt_rp_r_nofrp s0)) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_isrpin r0 s0 (l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) r s (l_e_st_eq_landau_n_rt_rp_r_nofrp s0) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) i) k j)))))))). Time Defined. (* constant 4043 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t30 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_amone l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_cut) => (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp y)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t29 r r x y t u (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)))))). Time Defined. (* constant 4044 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t31 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_negd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_negex r a0 a0ir n1)))). Time Defined. (* constant 4045 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_nr : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_cut)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_rpofnd a0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t31 r a0 a0ir n1))))). Time Defined. (* constant 4046 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t32 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp (l_e_st_eq_landau_n_rt_rp_r_ivr1_nr r a0 a0ir n1)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_isin r (l_e_st_eq_landau_n_rt_rp_r_nofrp (l_e_st_eq_landau_n_rt_rp_r_ivr1_nr r a0 a0ir n1)) a0 (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_r_ivr1_nr r a0 a0ir n1)) a0ir (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_ndofrp (l_e_st_eq_landau_n_rt_rp_r_ivr1_nr r a0 a0ir n1))) (l_e_st_eq_landau_n_rt_rp_eqndrp1 a0 (l_e_st_eq_landau_n_rt_rp_r_ivr1_t31 r a0 a0ir n1)))))). Time Defined. (* constant 4047 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t33 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_some (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_somei l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_nr r a0 a0ir n1) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t32 r a0 a0ir n1))))). Time Defined. (* constant 4048 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t34 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_some (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_some (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t33 r x t n)))). Time Defined. (* constant 4049 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr1_t35 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_one (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_onei l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t30 r) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t34 r n))). Time Defined. (* constant 4050 *) Definition l_e_st_eq_landau_n_rt_rp_r_rpofn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_cut)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_ind l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t35 r n))). Time Defined. (* constant 4051 *) Definition l_e_st_eq_landau_n_rt_rp_r_isnrp1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp (l_e_st_eq_landau_n_rt_rp_r_rpofn r n)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_oneax l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_nofrp x)) (l_e_st_eq_landau_n_rt_rp_r_ivr1_t35 r n))). Time Defined. (* constant 4052 *) Definition l_e_st_eq_landau_n_rt_rp_r_isnrp2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_nofrp (l_e_st_eq_landau_n_rt_rp_r_rpofn r n)) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_nofrp (l_e_st_eq_landau_n_rt_rp_r_rpofn r n)) (l_e_st_eq_landau_n_rt_rp_r_isnrp1 r n))). Time Defined. (* constant 4053 *) Definition l_e_st_eq_landau_n_rt_rp_r_isnerp : (forall (r1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r1), (forall (s1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg s1), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r1 s1), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_r_rpofn r1 n) (l_e_st_eq_landau_n_rt_rp_r_rpofn s1 m)))))). exact (fun (r1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r1) => (fun (s1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg s1) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r1 s1) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t29 r1 s1 (l_e_st_eq_landau_n_rt_rp_r_rpofn r1 n) (l_e_st_eq_landau_n_rt_rp_r_rpofn s1 m) (l_e_st_eq_landau_n_rt_rp_r_isnrp1 r1 n) (l_e_st_eq_landau_n_rt_rp_r_isnrp1 s1 m) i))))). Time Defined. (* constant 4054 *) Definition l_e_st_eq_landau_n_rt_rp_r_isnirp : (forall (r1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r1), (forall (s1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg s1), (forall (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_r_rpofn r1 n) (l_e_st_eq_landau_n_rt_rp_r_rpofn s1 m)), l_e_st_eq_landau_n_rt_rp_r_is r1 s1))))). exact (fun (r1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r1) => (fun (s1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg s1) => (fun (i:l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_r_rpofn r1 n) (l_e_st_eq_landau_n_rt_rp_r_rpofn s1 m)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real r1 (l_e_st_eq_landau_n_rt_rp_r_nofrp (l_e_st_eq_landau_n_rt_rp_r_rpofn r1 n)) (l_e_st_eq_landau_n_rt_rp_r_nofrp (l_e_st_eq_landau_n_rt_rp_r_rpofn s1 m)) s1 (l_e_st_eq_landau_n_rt_rp_r_isnrp1 r1 n) (l_e_st_eq_landau_n_rt_rp_r_isrpen (l_e_st_eq_landau_n_rt_rp_r_rpofn r1 n) (l_e_st_eq_landau_n_rt_rp_r_rpofn s1 m) i) (l_e_st_eq_landau_n_rt_rp_r_isnrp2 s1 m)))))). Time Defined. (* constant 4055 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrpn1 : (forall (r0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is r0 (l_e_st_eq_landau_n_rt_rp_r_rpofn (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_negi r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_ivr1_t29 (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) r0 (l_e_st_eq_landau_n_rt_rp_r_rpofn (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_negi r0)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_nofrp r0)) (l_e_st_eq_landau_n_rt_rp_r_isnrp1 (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_negi r0)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_nofrp r0))). Time Defined. (* constant 4056 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrpn2 : (forall (r0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_is (l_e_st_eq_landau_n_rt_rp_r_rpofn (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_negi r0)) r0). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_e_symis l_e_st_eq_landau_n_rt_cut r0 (l_e_st_eq_landau_n_rt_rp_r_rpofn (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_r_negi r0)) (l_e_st_eq_landau_n_rt_rp_r_isrpn1 r0)). Time Defined. (* constant 4057 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz163 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is r r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r). Time Defined. (* constant 4058 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz164 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r s i))). Time Defined. (* constant 4059 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz165 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is s t), l_e_st_eq_landau_n_rt_rp_r_is r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is s t) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real r s t i j))))). Time Defined. (* constant 4060 *) Definition l_e_st_eq_landau_n_rt_rp_r_absdr : (forall (x:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_absd x)). Time Defined. (* constant 4061 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_absdr a) (l_e_st_eq_landau_n_rt_rp_r_absdr b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_absd a) (l_e_st_eq_landau_n_rt_rp_absd b) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_absd a)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_absd b)) (l_e_st_eq_landau_n_rt_rp_eqabsd a b e)))). Time Defined. (* constant 4062 *) Definition l_e_st_eq_landau_n_rt_rp_r_fabsdr : l_e_st_eq_landau_n_rt_rp_r_fixf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_absdr. exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_eq x y) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t1 x y t))). Time Defined. (* constant 4063 *) Definition l_e_st_eq_landau_n_rt_rp_r_abs : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_indreal l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_absdr l_e_st_eq_landau_n_rt_rp_r_fabsdr r). Time Defined. (* constant 4064 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_absd a0)) (l_e_st_eq_landau_n_rt_rp_r_abs r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_isindreal l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_absdr l_e_st_eq_landau_n_rt_rp_r_fabsdr r a0 a0ir))). Time Defined. (* constant 4065 *) Definition l_e_st_eq_landau_n_rt_rp_r_aica : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_r_class (l_e_st_eq_landau_n_rt_rp_r_abs r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_r_class x)) (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_absd a0)) (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_absd a0)) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t2 r a0 a0ir)))). Time Defined. (* constant 4066 *) Definition l_e_st_eq_landau_n_rt_rp_r_isabs : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_abs x) r s i))). Time Defined. (* constant 4067 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_satzd166a a0 (l_e_st_eq_landau_n_rt_rp_r_posex r a0 a0ir p))))). Time Defined. (* constant 4068 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_r_aica r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_2r166_t1 r a0 a0ir p))))). Time Defined. (* constant 4069 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz166a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_2r166_t2 r x t p)))). Time Defined. (* constant 4070 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_absd a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_satzd166b a0 (l_e_st_eq_landau_n_rt_rp_r_negex r a0 a0ir n))))). Time Defined. (* constant 4071 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_r_aica r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_2r166_t3 r a0 a0ir n))))). Time Defined. (* constant 4072 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz166b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_2r166_t4 r x t n)))). Time Defined. (* constant 4073 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)), l_e_st_eq_landau_n_rt_rp_eq a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) => l_e_st_eq_landau_n_rt_rp_satzd166c a1 b1 (l_e_st_eq_landau_n_rt_rp_r_posex r a1 a1ir p) (l_e_st_eq_landau_n_rt_rp_r_posex s b1 b1is q) (l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1) (l_e_st_eq_landau_n_rt_rp_r_aica r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_aica s b1 b1is) i)))))))))). Time Defined. (* constant 4074 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)), l_e_st_eq_landau_n_rt_rp_r_is r s))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) => l_e_st_eq_landau_n_rt_rp_r_isin r s a1 b1 a1ir b1is (l_e_st_eq_landau_n_rt_rp_r_2r166_t5 r s a1 b1 a1ir b1is p q i)))))))))). Time Defined. (* constant 4075 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz166c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)), l_e_st_eq_landau_n_rt_rp_r_is r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_2r166_t6 r s x y t u p q i))))))))). Time Defined. (* constant 4076 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)), l_e_st_eq_landau_n_rt_rp_eq a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) => l_e_st_eq_landau_n_rt_rp_satzd166d a1 b1 (l_e_st_eq_landau_n_rt_rp_r_negex r a1 a1ir n) (l_e_st_eq_landau_n_rt_rp_r_negex s b1 b1is o) (l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1) (l_e_st_eq_landau_n_rt_rp_r_aica r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_aica s b1 b1is) i)))))))))). Time Defined. (* constant 4077 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)), l_e_st_eq_landau_n_rt_rp_r_is r s))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) => l_e_st_eq_landau_n_rt_rp_r_isin r s a1 b1 a1ir b1is (l_e_st_eq_landau_n_rt_rp_r_2r166_t7 r s a1 b1 a1ir b1is n o i)))))))))). Time Defined. (* constant 4078 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz166d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)), l_e_st_eq_landau_n_rt_rp_r_is r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_2r166_t8 r s x y t u n o i))))))))). Time Defined. (* constant 4079 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz166e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_rapp r (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_satz166a r t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs r)) n) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_satz166b r t))). Time Defined. (* constant 4080 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_absd a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_satzd166f a0 (l_e_st_eq_landau_n_rt_rp_r_0ex r a0 a0ir i))))). Time Defined. (* constant 4081 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r166_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_0in (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_r_aica r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_2r166_t9 r a0 a0ir i))))). Time Defined. (* constant 4082 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz166f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_2r166_t10 r x t i)))). Time Defined. (* constant 4083 *) Definition l_e_st_eq_landau_n_rt_rp_r_more : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_mored x y))))). Time Defined. (* constant 4084 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_propm : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), Prop)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_mored a1 b1))))). Time Defined. (* constant 4085 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a1 b1), l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a1 b1) => l_and3i (l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_mored a1 b1) a1ir b1is m))))))). Time Defined. (* constant 4086 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a1 b1), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a1 x)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a1 b1) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a1 x) b1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t3 r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4087 *) Definition l_e_st_eq_landau_n_rt_rp_r_morein : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a1 b1), l_e_st_eq_landau_n_rt_rp_r_more r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a1 b1) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s x y)) a1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t4 r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4088 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a b), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a b) => l_and3e1 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_mored a b) p2))))))))))). Time Defined. (* constant 4089 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a b), l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a b) => l_and3e2 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_mored a b) p2))))))))))). Time Defined. (* constant 4090 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a b), l_e_st_eq_landau_n_rt_rp_mored a b))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a b) => l_and3e3 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_mored a b) p2))))))))))). Time Defined. (* constant 4091 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a b), l_e_st_eq_landau_n_rt_rp_mored a1 b1))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a b) => l_e_st_eq_landau_n_rt_rp_eqmored12 a a1 b b1 (l_e_st_eq_landau_n_rt_rp_r_isex r r a a1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t5 r s a1 b1 a1ir b1is m a sa b p2) a1ir (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)) (l_e_st_eq_landau_n_rt_rp_r_isex s s b b1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t6 r s a1 b1 a1ir b1is m a sa b p2) b1is (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real s)) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t7 r s a1 b1 a1ir b1is m a sa b p2)))))))))))). Time Defined. (* constant 4092 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)), l_e_st_eq_landau_n_rt_rp_mored a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x)) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x) sa (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s a x) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t8 r s a1 b1 a1ir b1is m a sa x t))))))))))). Time Defined. (* constant 4093 *) Definition l_e_st_eq_landau_n_rt_rp_r_moreex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_mored a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s x y)) m (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propm r s x y)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t9 r s a1 b1 a1ir b1is m x t))))))))). Time Defined. (* constant 4094 *) Definition l_e_st_eq_landau_n_rt_rp_r_less : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_lessd x y))))). Time Defined. (* constant 4095 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_propl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), Prop)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1))))). Time Defined. (* constant 4096 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a1 b1), l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a1 b1) => l_and3i (l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) a1ir b1is l))))))). Time Defined. (* constant 4097 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a1 b1), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a1 x)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a1 b1) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a1 x) b1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t10 r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4098 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a1 b1), l_e_st_eq_landau_n_rt_rp_r_less r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a1 b1) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s x y)) a1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t11 r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4099 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a b), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a b) => l_and3e1 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_lessd a b) p2))))))))))). Time Defined. (* constant 4100 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t13 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a b), l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a b) => l_and3e2 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_lessd a b) p2))))))))))). Time Defined. (* constant 4101 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t14 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a b), l_e_st_eq_landau_n_rt_rp_lessd a b))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a b) => l_and3e3 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)) (l_e_st_eq_landau_n_rt_rp_lessd a b) p2))))))))))). Time Defined. (* constant 4102 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t15 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a b), l_e_st_eq_landau_n_rt_rp_lessd a1 b1))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a b) => l_e_st_eq_landau_n_rt_rp_eqlessd12 a a1 b b1 (l_e_st_eq_landau_n_rt_rp_r_isex r r a a1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t12 r s a1 b1 a1ir b1is l a sa b p2) a1ir (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)) (l_e_st_eq_landau_n_rt_rp_r_isex s s b b1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t13 r s a1 b1 a1ir b1is l a sa b p2) b1is (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real s)) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t14 r s a1 b1 a1ir b1is l a sa b p2)))))))))))). Time Defined. (* constant 4103 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t16 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)), l_e_st_eq_landau_n_rt_rp_lessd a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (sa:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x)) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x) sa (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s a x) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t15 r s a1 b1 a1ir b1is l a sa x t))))))))))). Time Defined. (* constant 4104 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_lessd a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s x y)) l (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_some l_e_st_eq_landau_n_rt_rp_dif (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_ivr2_propl r s x y)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t16 r s a1 b1 a1ir b1is l x t))))))))). Time Defined. (* constant 4105 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismore1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r t), l_e_st_eq_landau_n_rt_rp_r_more s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r t) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_more x t) r s m i))))). Time Defined. (* constant 4106 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismore2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more t r), l_e_st_eq_landau_n_rt_rp_r_more t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more t r) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_more t x) r s m i))))). Time Defined. (* constant 4107 *) Definition l_e_st_eq_landau_n_rt_rp_r_isless1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r t), l_e_st_eq_landau_n_rt_rp_r_less s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r t) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_less x t) r s l i))))). Time Defined. (* constant 4108 *) Definition l_e_st_eq_landau_n_rt_rp_r_isless2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less t r), l_e_st_eq_landau_n_rt_rp_r_less t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less t r) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_less t x) r s l i))))). Time Defined. (* constant 4109 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismore12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is t u), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r t), l_e_st_eq_landau_n_rt_rp_r_more s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is t u) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r t) => l_e_st_eq_landau_n_rt_rp_r_ismore2 t u s j (l_e_st_eq_landau_n_rt_rp_r_ismore1 r s t i m)))))))). Time Defined. (* constant 4110 *) Definition l_e_st_eq_landau_n_rt_rp_r_isless12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is t u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r t), l_e_st_eq_landau_n_rt_rp_r_less s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is t u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r t) => l_e_st_eq_landau_n_rt_rp_r_isless2 t u s j (l_e_st_eq_landau_n_rt_rp_r_isless1 r s t i l)))))))). Time Defined. (* constant 4111 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t17 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_lessd b1 a1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_lemmad5 a1 b1 (l_e_st_eq_landau_n_rt_rp_r_moreex r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4112 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t18 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_less s r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_lessin s r b1 a1 b1is a1ir (l_e_st_eq_landau_n_rt_rp_r_ivr2_t17 r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4113 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_less s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_less s r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t18 r s x y t u m))))))). Time Defined. (* constant 4114 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t19 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_mored b1 a1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_lemmad6 a1 b1 (l_e_st_eq_landau_n_rt_rp_r_lessex r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4115 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t20 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_more s r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_morein s r b1 a1 b1is a1ir (l_e_st_eq_landau_n_rt_rp_r_ivr2_t19 r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4116 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_more s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_more s r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t20 r s x y t u l))))))). Time Defined. (* constant 4117 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_or3 (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_satzd167a a1 b1)))))). Time Defined. (* constant 4118 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a1 b1), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a1 b1) => l_or3i1 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_isin r s a1 b1 a1ir b1is e)))))))). Time Defined. (* constant 4119 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_mored a1 b1), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_mored a1 b1) => l_or3i2 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_morein r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4120 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_lessd a1 b1), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_lessd a1 b1) => l_or3i3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_lessin r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4121 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_or3app (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) (l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s)) (l_e_st_eq_landau_n_rt_rp_r_2r167_t1 r s a1 b1 a1ir b1is) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a1 b1) => l_e_st_eq_landau_n_rt_rp_r_2r167_t2 r s a1 b1 a1ir b1is t) (fun (t:l_e_st_eq_landau_n_rt_rp_mored a1 b1) => l_e_st_eq_landau_n_rt_rp_r_2r167_t3 r s a1 b1 a1ir b1is t) (fun (t:l_e_st_eq_landau_n_rt_rp_lessd a1 b1) => l_e_st_eq_landau_n_rt_rp_r_2r167_t4 r s a1 b1 a1ir b1is t))))))). Time Defined. (* constant 4122 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_ec3 (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_satzd167b a1 b1)))))). Time Defined. (* constant 4123 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_not (l_e_st_eq_landau_n_rt_rp_r_more r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (l_ec3e12 (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_2r167_t6 r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_isex r s a1 b1 a1ir b1is i)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_moreex r s a1 b1 a1ir b1is t)))))))). Time Defined. (* constant 4124 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_not (l_e_st_eq_landau_n_rt_rp_r_less r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) (l_ec3e23 (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_2r167_t6 r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_moreex r s a1 b1 a1ir b1is m)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_lessex r s a1 b1 a1ir b1is t)))))))). Time Defined. (* constant 4125 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_not (l_e_st_eq_landau_n_rt_rp_r_is r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_ec3e31 (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_2r167_t6 r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_lessex r s a1 b1 a1ir b1is l)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_isex r s a1 b1 a1ir b1is t)))))))). Time Defined. (* constant 4126 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_ec3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_ec3_th6 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_ec_th1 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_2r167_t7 r s a1 b1 a1ir b1is t)) (l_ec_th1 (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (fun (t:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_2r167_t8 r s a1 b1 a1ir b1is t)) (l_ec_th1 (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_2r167_t9 r s a1 b1 a1ir b1is t)))))))). Time Defined. (* constant 4127 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r167_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_orec3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_orec3i (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_2r167_t5 r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_2r167_t10 r s a1 b1 a1ir b1is))))))). Time Defined. (* constant 4128 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_orec3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_orec3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_2r167_t11 r s x y t u)))))). Time Defined. (* constant 4129 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_or3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_orec3e1 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_satz167 r s))). Time Defined. (* constant 4130 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_ec3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_orec3e2 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_satz167 r s))). Time Defined. (* constant 4131 *) Definition l_e_st_eq_landau_n_rt_rp_r_moreis : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s))). Time Defined. (* constant 4132 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessis : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_is r s))). Time Defined. (* constant 4133 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz168a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), l_e_st_eq_landau_n_rt_rp_r_lessis s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_less s r) (l_e_st_eq_landau_n_rt_rp_r_is s r) m (fun (t:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_lemma1 r s t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r s t)))). Time Defined. (* constant 4134 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz168b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), l_e_st_eq_landau_n_rt_rp_r_moreis s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more s r) (l_e_st_eq_landau_n_rt_rp_r_is s r) l (fun (t:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_lemma2 r s t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r s t)))). Time Defined. (* constant 4135 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismoreis1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r t), l_e_st_eq_landau_n_rt_rp_r_moreis s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r t) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_moreis x t) r s m i))))). Time Defined. (* constant 4136 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismoreis2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis t r), l_e_st_eq_landau_n_rt_rp_r_moreis t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis t r) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_moreis t x) r s m i))))). Time Defined. (* constant 4137 *) Definition l_e_st_eq_landau_n_rt_rp_r_islessis1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r t), l_e_st_eq_landau_n_rt_rp_r_lessis s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r t) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_lessis x t) r s l i))))). Time Defined. (* constant 4138 *) Definition l_e_st_eq_landau_n_rt_rp_r_islessis2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis t r), l_e_st_eq_landau_n_rt_rp_r_lessis t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis t r) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_lessis t x) r s l i))))). Time Defined. (* constant 4139 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismoreis12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is t u), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r t), l_e_st_eq_landau_n_rt_rp_r_moreis s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is t u) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r t) => l_e_st_eq_landau_n_rt_rp_r_ismoreis2 t u s j (l_e_st_eq_landau_n_rt_rp_r_ismoreis1 r s t i m)))))))). Time Defined. (* constant 4140 *) Definition l_e_st_eq_landau_n_rt_rp_r_islessis12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is t u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r t), l_e_st_eq_landau_n_rt_rp_r_lessis s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is t u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r t) => l_e_st_eq_landau_n_rt_rp_r_islessis2 t u s j (l_e_st_eq_landau_n_rt_rp_r_islessis1 r s t i l)))))))). Time Defined. (* constant 4141 *) Definition l_e_st_eq_landau_n_rt_rp_r_moreisi1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_moreis r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) m))). Time Defined. (* constant 4142 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessisi1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_lessis r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) l))). Time Defined. (* constant 4143 *) Definition l_e_st_eq_landau_n_rt_rp_r_moreisi2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_moreis r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) i))). Time Defined. (* constant 4144 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessisi2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_lessis r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) i))). Time Defined. (* constant 4145 *) Definition l_e_st_eq_landau_n_rt_rp_r_moreisin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_moreq a1 b1), l_e_st_eq_landau_n_rt_rp_r_moreis r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_moreq a1 b1) => l_orapp (l_e_st_eq_landau_n_rt_rp_mored a1 b1) (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_r_moreis r s) m (fun (t:l_e_st_eq_landau_n_rt_rp_mored a1 b1) => l_e_st_eq_landau_n_rt_rp_r_moreisi1 r s (l_e_st_eq_landau_n_rt_rp_r_morein r s a1 b1 a1ir b1is t)) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a1 b1) => l_e_st_eq_landau_n_rt_rp_r_moreisi2 r s (l_e_st_eq_landau_n_rt_rp_r_isin r s a1 b1 a1ir b1is t))))))))). Time Defined. (* constant 4146 *) Definition l_e_st_eq_landau_n_rt_rp_r_moreisex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), l_e_st_eq_landau_n_rt_rp_moreq a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => l_orapp (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_moreq a1 b1) m (fun (t:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_moreqi1 a1 b1 (l_e_st_eq_landau_n_rt_rp_r_moreex r s a1 b1 a1ir b1is t)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_moreqi2 a1 b1 (l_e_st_eq_landau_n_rt_rp_r_isex r s a1 b1 a1ir b1is t))))))))). Time Defined. (* constant 4147 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessisin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_lesseq a1 b1), l_e_st_eq_landau_n_rt_rp_r_lessis r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_lesseq a1 b1) => l_orapp (l_e_st_eq_landau_n_rt_rp_lessd a1 b1) (l_e_st_eq_landau_n_rt_rp_eq a1 b1) (l_e_st_eq_landau_n_rt_rp_r_lessis r s) l (fun (t:l_e_st_eq_landau_n_rt_rp_lessd a1 b1) => l_e_st_eq_landau_n_rt_rp_r_lessisi1 r s (l_e_st_eq_landau_n_rt_rp_r_lessin r s a1 b1 a1ir b1is t)) (fun (t:l_e_st_eq_landau_n_rt_rp_eq a1 b1) => l_e_st_eq_landau_n_rt_rp_r_lessisi2 r s (l_e_st_eq_landau_n_rt_rp_r_isin r s a1 b1 a1ir b1is t))))))))). Time Defined. (* constant 4148 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessisex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), l_e_st_eq_landau_n_rt_rp_lesseq a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => l_orapp (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_lesseq a1 b1) l (fun (t:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_lesseqi1 a1 b1 (l_e_st_eq_landau_n_rt_rp_r_lessex r s a1 b1 a1ir b1is t)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_lesseqi2 a1 b1 (l_e_st_eq_landau_n_rt_rp_r_isex r s a1 b1 a1ir b1is t))))))))). Time Defined. (* constant 4149 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), l_not (l_e_st_eq_landau_n_rt_rp_r_less r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => l_ec3_th7 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_satz167b r s) (l_comor (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) m)))). Time Defined. (* constant 4150 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), l_not (l_e_st_eq_landau_n_rt_rp_r_more r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => l_ec3_th9 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_satz167b r s) l))). Time Defined. (* constant 4151 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_more r s)), l_e_st_eq_landau_n_rt_rp_r_lessis r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_more r s)) => l_or3_th2 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_satz167a r s) n))). Time Defined. (* constant 4152 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_less r s)), l_e_st_eq_landau_n_rt_rp_r_moreis r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_less r s)) => l_comor (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_or3_th3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_satz167a r s) n)))). Time Defined. (* constant 4153 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167g : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_not (l_e_st_eq_landau_n_rt_rp_r_lessis r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_lessis r s) (l_not (l_e_st_eq_landau_n_rt_rp_r_more r s)) (l_weli (l_e_st_eq_landau_n_rt_rp_r_more r s) m) (fun (t:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => l_e_st_eq_landau_n_rt_rp_r_satz167d r s t)))). Time Defined. (* constant 4154 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167h : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_not (l_e_st_eq_landau_n_rt_rp_r_moreis r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_moreis r s) (l_not (l_e_st_eq_landau_n_rt_rp_r_less r s)) (l_weli (l_e_st_eq_landau_n_rt_rp_r_less r s) l) (fun (t:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => l_e_st_eq_landau_n_rt_rp_r_satz167c r s t)))). Time Defined. (* constant 4155 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167j : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_moreis r s)), l_e_st_eq_landau_n_rt_rp_r_less r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_moreis r s)) => l_or3e3 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_satz167a r s) (l_or_th5 (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) n) (l_or_th4 (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) n)))). Time Defined. (* constant 4156 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz167k : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_lessis r s)), l_e_st_eq_landau_n_rt_rp_r_more r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_lessis r s)) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_satz167a r s) (l_or_th4 (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) n) (l_or_th5 (l_e_st_eq_landau_n_rt_rp_r_less r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) n)))). Time Defined. (* constant 4157 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r169_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_mored a b)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_satzd169a a b (l_e_st_eq_landau_n_rt_rp_r_0ex l_e_st_eq_landau_n_rt_rp_r_0 b bi0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_posex r a air p))))))). Time Defined. (* constant 4158 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r169_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_morein r l_e_st_eq_landau_n_rt_rp_r_0 a b air bi0 (l_e_st_eq_landau_n_rt_rp_r_2r169_t1 r p a b air bi0))))))). Time Defined. (* constant 4159 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz169a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_2r169_t2 r p x y t u)))))). Time Defined. (* constant 4160 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r169_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_posd a)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_satzd169b a b (l_e_st_eq_landau_n_rt_rp_r_0ex l_e_st_eq_landau_n_rt_rp_r_0 b bi0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_moreex r l_e_st_eq_landau_n_rt_rp_r_0 a b air bi0 m))))))). Time Defined. (* constant 4161 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r169_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_r_pos r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_posin r a air (l_e_st_eq_landau_n_rt_rp_r_2r169_t3 r m a b air bi0))))))). Time Defined. (* constant 4162 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz169b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_pos r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pos r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_2r169_t4 r m x y t u)))))). Time Defined. (* constant 4163 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r169_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_lessd a b)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_satzd169c a b (l_e_st_eq_landau_n_rt_rp_r_0ex l_e_st_eq_landau_n_rt_rp_r_0 b bi0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_negex r a air n))))))). Time Defined. (* constant 4164 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r169_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_lessin r l_e_st_eq_landau_n_rt_rp_r_0 a b air bi0 (l_e_st_eq_landau_n_rt_rp_r_2r169_t5 r n a b air bi0))))))). Time Defined. (* constant 4165 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz169c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_2r169_t6 r n x y t u)))))). Time Defined. (* constant 4166 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r169_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_negd a)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_satzd169d a b (l_e_st_eq_landau_n_rt_rp_r_0ex l_e_st_eq_landau_n_rt_rp_r_0 b bi0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_lessex r l_e_st_eq_landau_n_rt_rp_r_0 a b air bi0 l))))))). Time Defined. (* constant 4167 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r169_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_r_neg r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_negin r a air (l_e_st_eq_landau_n_rt_rp_r_2r169_t7 r l a b air bi0))))))). Time Defined. (* constant 4168 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz169d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_neg r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_neg r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_2r169_t8 r l x y t u)))))). Time Defined. (* constant 4169 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r170_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_absd a) b))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_satzd170 a b (l_e_st_eq_landau_n_rt_rp_r_0ex l_e_st_eq_landau_n_rt_rp_r_0 b bi0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))))). Time Defined. (* constant 4170 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r170_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bi0:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_moreisin (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_absd a) b (l_e_st_eq_landau_n_rt_rp_r_aica r a air) bi0 (l_e_st_eq_landau_n_rt_rp_r_2r170_t1 r a b air bi0)))))). Time Defined. (* constant 4171 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz170 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class l_e_st_eq_landau_n_rt_rp_r_0)) => l_e_st_eq_landau_n_rt_rp_r_2r170_t2 r x y t u))))). Time Defined. (* constant 4172 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz170a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_abs r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_abs r)) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz167c (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satz170 r)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_abs r)) => l_e_st_eq_landau_n_rt_rp_r_satz169c (l_e_st_eq_landau_n_rt_rp_r_abs r) t)). Time Defined. (* constant 4173 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r171_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less s t), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bis:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_lessd a c))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less s t) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bis:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_satzd171 a b c (l_e_st_eq_landau_n_rt_rp_r_lessex r s a b air bis l) (l_e_st_eq_landau_n_rt_rp_r_lessex s t b c bis cit k)))))))))))). Time Defined. (* constant 4174 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r171_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less s t), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (bis:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_r_less r t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less s t) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (bis:l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_lessin r t a c air cit (l_e_st_eq_landau_n_rt_rp_r_2r171_t1 r s t l k a b c air bis cit)))))))))))). Time Defined. (* constant 4175 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz171 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less s t), l_e_st_eq_landau_n_rt_rp_r_less r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less s t) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_less r t) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_2r171_t2 r s t l k x y z w u v))))))))))). Time Defined. (* constant 4176 *) Definition l_e_st_eq_landau_n_rt_rp_r_trless : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less s t), l_e_st_eq_landau_n_rt_rp_r_less r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less s t) => l_e_st_eq_landau_n_rt_rp_r_satz171 r s t l k))))). Time Defined. (* constant 4177 *) Definition l_e_st_eq_landau_n_rt_rp_r_trmore : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_more s t), l_e_st_eq_landau_n_rt_rp_r_more r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_more s t) => l_e_st_eq_landau_n_rt_rp_r_lemma2 t r (l_e_st_eq_landau_n_rt_rp_r_trless t s r (l_e_st_eq_landau_n_rt_rp_r_lemma1 s t n) (l_e_st_eq_landau_n_rt_rp_r_lemma1 r s m))))))). Time Defined. (* constant 4178 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r172_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less s t), l_e_st_eq_landau_n_rt_rp_lessd a2 c2))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less s t) => l_e_st_eq_landau_n_rt_rp_satzd172a a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_lessisex r s a2 b2 a2ir b2is l) (l_e_st_eq_landau_n_rt_rp_r_lessex s t b2 c2 b2is c2it k)))))))))))). Time Defined. (* constant 4179 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r172_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less s t), l_e_st_eq_landau_n_rt_rp_r_less r t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less s t) => l_e_st_eq_landau_n_rt_rp_r_lessin r t a2 c2 a2ir c2it (l_e_st_eq_landau_n_rt_rp_r_2r172_t1 r s t a2 b2 c2 a2ir b2is c2it l k)))))))))))). Time Defined. (* constant 4180 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz172a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less s t), l_e_st_eq_landau_n_rt_rp_r_less r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less s t) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_less r t) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_2r172_t2 r s t x y z u v w l k))))))))))). Time Defined. (* constant 4181 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r172_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t), l_e_st_eq_landau_n_rt_rp_lessd a2 c2))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t) => l_e_st_eq_landau_n_rt_rp_satzd172b a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_lessex r s a2 b2 a2ir b2is l) (l_e_st_eq_landau_n_rt_rp_r_lessisex s t b2 c2 b2is c2it k)))))))))))). Time Defined. (* constant 4182 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r172_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t), l_e_st_eq_landau_n_rt_rp_r_less r t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t) => l_e_st_eq_landau_n_rt_rp_r_lessin r t a2 c2 a2ir c2it (l_e_st_eq_landau_n_rt_rp_r_2r172_t3 r s t a2 b2 c2 a2ir b2is c2it l k)))))))))))). Time Defined. (* constant 4183 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz172b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t), l_e_st_eq_landau_n_rt_rp_r_less r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_less r t) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_2r172_t4 r s t x y z u v w l k))))))))))). Time Defined. (* constant 4184 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz172c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_more s t), l_e_st_eq_landau_n_rt_rp_r_more r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_more s t) => l_e_st_eq_landau_n_rt_rp_r_lemma2 t r (l_e_st_eq_landau_n_rt_rp_r_satz172b t s r (l_e_st_eq_landau_n_rt_rp_r_lemma1 s t n) (l_e_st_eq_landau_n_rt_rp_r_satz168a r s m))))))). Time Defined. (* constant 4185 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz172d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis s t), l_e_st_eq_landau_n_rt_rp_r_more r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis s t) => l_e_st_eq_landau_n_rt_rp_r_lemma2 t r (l_e_st_eq_landau_n_rt_rp_r_satz172a t s r (l_e_st_eq_landau_n_rt_rp_r_satz168a s t n) (l_e_st_eq_landau_n_rt_rp_r_lemma1 r s m))))))). Time Defined. (* constant 4186 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r173_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t), l_e_st_eq_landau_n_rt_rp_lesseq a2 c2))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t) => l_e_st_eq_landau_n_rt_rp_satzd173 a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_lessisex r s a2 b2 a2ir b2is l) (l_e_st_eq_landau_n_rt_rp_r_lessisex s t b2 c2 b2is c2it k)))))))))))). Time Defined. (* constant 4187 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r173_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t), l_e_st_eq_landau_n_rt_rp_r_lessis r t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t) => l_e_st_eq_landau_n_rt_rp_r_lessisin r t a2 c2 a2ir c2it (l_e_st_eq_landau_n_rt_rp_r_2r173_t1 r s t a2 b2 c2 a2ir b2is c2it l k)))))))))))). Time Defined. (* constant 4188 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz173 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t), l_e_st_eq_landau_n_rt_rp_r_lessis r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_lessis r t) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_2r173_t2 r s t x y z u v w l k))))))))))). Time Defined. (* constant 4189 *) Definition l_e_st_eq_landau_n_rt_rp_r_trlessis : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t), l_e_st_eq_landau_n_rt_rp_r_lessis r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis s t) => l_e_st_eq_landau_n_rt_rp_r_satz173 r s t l k))))). Time Defined. (* constant 4190 *) Definition l_e_st_eq_landau_n_rt_rp_r_trmoreis : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis s t), l_e_st_eq_landau_n_rt_rp_r_moreis r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis s t) => l_e_st_eq_landau_n_rt_rp_r_satz168b t r (l_e_st_eq_landau_n_rt_rp_r_trlessis t s r (l_e_st_eq_landau_n_rt_rp_r_satz168a s t n) (l_e_st_eq_landau_n_rt_rp_r_satz168a r s m))))))). Time Defined. (* constant 4191 *) Definition l_e_st_eq_landau_n_rt_rp_r_ratrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd x))). Time Defined. (* constant 4192 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t21 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (r1:l_e_st_eq_landau_n_rt_rp_ratd a0), l_and (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (r1:l_e_st_eq_landau_n_rt_rp_ratd a0) => l_andi (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a0) a0ir r1)))). Time Defined. (* constant 4193 *) Definition l_e_st_eq_landau_n_rt_rp_r_ratrlin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (r1:l_e_st_eq_landau_n_rt_rp_ratd a0), l_e_st_eq_landau_n_rt_rp_r_ratrl r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (r1:l_e_st_eq_landau_n_rt_rp_ratd a0) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd x)) a0 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t21 r a0 a0ir r1))))). Time Defined. (* constant 4194 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t22 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (rr:l_e_st_eq_landau_n_rt_rp_r_ratrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a)), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (rr:l_e_st_eq_landau_n_rt_rp_r_ratrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a) b)))))). Time Defined. (* constant 4195 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t23 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (rr:l_e_st_eq_landau_n_rt_rp_r_ratrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a)), l_e_st_eq_landau_n_rt_rp_ratd a)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (rr:l_e_st_eq_landau_n_rt_rp_r_ratrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a) b)))))). Time Defined. (* constant 4196 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t24 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (rr:l_e_st_eq_landau_n_rt_rp_r_ratrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a)), l_e_st_eq_landau_n_rt_rp_ratd a0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (rr:l_e_st_eq_landau_n_rt_rp_r_ratrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd a)) => l_e_st_eq_landau_n_rt_rp_eqratd a a0 (l_e_st_eq_landau_n_rt_rp_r_isex r r a a0 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t22 r a0 a0ir rr a b) a0ir (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t23 r a0 a0ir rr a b))))))). Time Defined. (* constant 4197 *) Definition l_e_st_eq_landau_n_rt_rp_r_ratrlex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (rr:l_e_st_eq_landau_n_rt_rp_r_ratrl r), l_e_st_eq_landau_n_rt_rp_ratd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (rr:l_e_st_eq_landau_n_rt_rp_r_ratrl r) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd x)) rr (l_e_st_eq_landau_n_rt_rp_ratd a0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_ratd x)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t24 r a0 a0ir rr x t)))))). Time Defined. (* constant 4198 *) Definition l_e_st_eq_landau_n_rt_rp_r_irratrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_not (l_e_st_eq_landau_n_rt_rp_r_ratrl r)). Time Defined. (* constant 4199 *) Definition l_e_st_eq_landau_n_rt_rp_r_remark2 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (rr:l_e_st_eq_landau_n_rt_rp_ratrp r0), l_e_st_eq_landau_n_rt_rp_r_ratrl (l_e_st_eq_landau_n_rt_rp_r_pofrp r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (rr:l_e_st_eq_landau_n_rt_rp_ratrp r0) => l_e_st_eq_landau_n_rt_rp_r_ratrlin (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_remark2a r0 rr))). Time Defined. (* constant 4200 *) Definition l_e_st_eq_landau_n_rt_rp_r_remark3 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (rr:l_e_st_eq_landau_n_rt_rp_ratrp r0), l_e_st_eq_landau_n_rt_rp_r_ratrl (l_e_st_eq_landau_n_rt_rp_r_nofrp r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (rr:l_e_st_eq_landau_n_rt_rp_ratrp r0) => l_e_st_eq_landau_n_rt_rp_r_ratrlin (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_ndofrp r0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_ndofrp r0)) (l_e_st_eq_landau_n_rt_rp_remark3a r0 rr))). Time Defined. (* constant 4201 *) Definition l_e_st_eq_landau_n_rt_rp_r_remark4 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (ir:l_e_st_eq_landau_n_rt_rp_irratrp r0), l_e_st_eq_landau_n_rt_rp_r_irratrl (l_e_st_eq_landau_n_rt_rp_r_pofrp r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (ir:l_e_st_eq_landau_n_rt_rp_irratrp r0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_ratrl (l_e_st_eq_landau_n_rt_rp_r_pofrp r0)) (l_e_st_eq_landau_n_rt_rp_ratd (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_remark4a r0 ir) (fun (t:l_e_st_eq_landau_n_rt_rp_r_ratrl (l_e_st_eq_landau_n_rt_rp_r_pofrp r0)) => l_e_st_eq_landau_n_rt_rp_r_ratrlex (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) t))). Time Defined. (* constant 4202 *) Definition l_e_st_eq_landau_n_rt_rp_r_remark5 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (ir:l_e_st_eq_landau_n_rt_rp_irratrp r0), l_e_st_eq_landau_n_rt_rp_r_irratrl (l_e_st_eq_landau_n_rt_rp_r_nofrp r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (ir:l_e_st_eq_landau_n_rt_rp_irratrp r0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_ratrl (l_e_st_eq_landau_n_rt_rp_r_nofrp r0)) (l_e_st_eq_landau_n_rt_rp_ratd (l_e_st_eq_landau_n_rt_rp_ndofrp r0)) (l_e_st_eq_landau_n_rt_rp_remark5a r0 ir) (fun (t:l_e_st_eq_landau_n_rt_rp_r_ratrl (l_e_st_eq_landau_n_rt_rp_r_nofrp r0)) => l_e_st_eq_landau_n_rt_rp_r_ratrlex (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_ndofrp r0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_ndofrp r0)) t))). Time Defined. (* constant 4203 *) Definition l_e_st_eq_landau_n_rt_rp_r_natrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd x))). Time Defined. (* constant 4204 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t25 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a0), l_and (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a0) => l_andi (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a0) a0ir n)))). Time Defined. (* constant 4205 *) Definition l_e_st_eq_landau_n_rt_rp_r_natrlin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_natd a0), l_e_st_eq_landau_n_rt_rp_r_natrl r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_natd a0) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd x)) a0 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t25 r a0 a0ir n))))). Time Defined. (* constant 4206 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t26 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a)), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a) b)))))). Time Defined. (* constant 4207 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t27 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a)), l_e_st_eq_landau_n_rt_rp_natd a)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a) b)))))). Time Defined. (* constant 4208 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t28 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a)), l_e_st_eq_landau_n_rt_rp_natd a0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd a)) => l_e_st_eq_landau_n_rt_rp_eqnatd a a0 (l_e_st_eq_landau_n_rt_rp_r_isex r r a a0 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t26 r a0 a0ir n a b) a0ir (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t27 r a0 a0ir n a b))))))). Time Defined. (* constant 4209 *) Definition l_e_st_eq_landau_n_rt_rp_r_natrlex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_natd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd x)) n (l_e_st_eq_landau_n_rt_rp_natd a0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_natd x)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t28 r a0 a0ir n x t)))))). Time Defined. (* constant 4210 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t29 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_posd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_natposd a0 (l_e_st_eq_landau_n_rt_rp_r_natrlex r a0 a0ir n))))). Time Defined. (* constant 4211 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t30 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_r_pos r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_r_posin r a0 a0ir (l_e_st_eq_landau_n_rt_rp_r_ivr2_t29 r a0 a0ir n))))). Time Defined. (* constant 4212 *) Definition l_e_st_eq_landau_n_rt_rp_r_natpos : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_r_pos r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_pos r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t30 r x t n)))). Time Defined. (* constant 4213 *) Definition l_e_st_eq_landau_n_rt_rp_r_rlofnt : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_pdofnt x)). Time Defined. (* constant 4214 *) Definition l_e_st_eq_landau_n_rt_rp_r_natrli : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_natrl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x)). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_natrlin (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_pdofnt x) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofnt x)) (l_e_st_eq_landau_n_rt_rp_natdi x)). Time Defined. (* constant 4215 *) Definition l_e_st_eq_landau_n_rt_rp_r_isnterl : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is x y), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is x y) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt u) x y i))). Time Defined. (* constant 4216 *) Definition l_e_st_eq_landau_n_rt_rp_r_isntirl : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)), l_e_st_eq_landau_n_is x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) => l_e_st_eq_landau_n_rt_rp_isntirp x y (l_e_st_eq_landau_n_rt_rp_r_isrpip (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y) i)))). Time Defined. (* constant 4217 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t31 : l_e_injective l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) => l_e_st_eq_landau_n_rt_rp_r_isntirl x y t))). Time Defined. (* constant 4218 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t32 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_posd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_natposd a0 (l_e_st_eq_landau_n_rt_rp_r_natrlex r a0 a0ir n))))). Time Defined. (* constant 4219 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_ap : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_cut)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_rpofpd a0 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t32 r a0 a0ir n))))). Time Defined. (* constant 4220 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t33 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_natrp (l_e_st_eq_landau_n_rt_rp_r_ivr2_ap r a0 a0ir n))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_natderp a0 (l_e_st_eq_landau_n_rt_rp_r_natrlex r a0 a0ir n))))). Time Defined. (* constant 4221 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_nat)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_ntofrp (l_e_st_eq_landau_n_rt_rp_r_ivr2_ap r a0 a0ir n) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t33 r a0 a0ir n))))). Time Defined. (* constant 4222 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t34 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_r_ivr2_ap r a0 a0ir n)) (l_e_st_eq_landau_n_rt_rp_pdofnt (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_isrpepd (l_e_st_eq_landau_n_rt_rp_r_ivr2_ap r a0 a0ir n) (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n)) (l_e_st_eq_landau_n_rt_rp_isrpnt1 (l_e_st_eq_landau_n_rt_rp_r_ivr2_ap r a0 a0ir n) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t33 r a0 a0ir n)))))). Time Defined. (* constant 4223 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t35 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_eq a0 (l_e_st_eq_landau_n_rt_rp_pdofnt (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_treq a0 (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_r_ivr2_ap r a0 a0ir n)) (l_e_st_eq_landau_n_rt_rp_pdofnt (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n)) (l_e_st_eq_landau_n_rt_rp_eqpdrp1 a0 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t32 r a0 a0ir n)) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t34 r a0 a0ir n))))). Time Defined. (* constant 4224 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t36 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_r_isin r (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n)) a0 (l_e_st_eq_landau_n_rt_rp_pdofnt (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n)) a0ir (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofnt (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n))) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t35 r a0 a0ir n))))). Time Defined. (* constant 4225 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t37 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_somei l_e_st_eq_landau_n_nat (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_rlofnt x)) (l_e_st_eq_landau_n_rt_rp_r_ivr2_x0 r a0 a0ir n) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t36 r a0 a0ir n))))). Time Defined. (* constant 4226 *) Definition l_e_st_eq_landau_n_rt_rp_r_natimage : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t37 r x t n)))). Time Defined. (* constant 4227 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t38 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) r), (forall (x:l_e_st_eq_landau_n_nat), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_rlofnt x)), l_e_st_eq_landau_n_rt_rp_r_natrl r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) r) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_rlofnt x)) => l_e_isp1 l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_natrl u) (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) r (l_e_st_eq_landau_n_rt_rp_r_natrli x) j)))). Time Defined. (* constant 4228 *) Definition l_e_st_eq_landau_n_rt_rp_r_imagenat : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) r), l_e_st_eq_landau_n_rt_rp_r_natrl r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) r) => l_someapp l_e_st_eq_landau_n_nat (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_rlofnt u)) i (l_e_st_eq_landau_n_rt_rp_r_natrl r) (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_rlofnt u)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t38 r i u v)))). Time Defined. (* constant 4229 *) Definition l_e_st_eq_landau_n_rt_rp_r_ntofrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_nat)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_soft l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) l_e_st_eq_landau_n_rt_rp_r_ivr2_t31 r (l_e_st_eq_landau_n_rt_rp_r_natimage r n))). Time Defined. (* constant 4230 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrlent : (forall (r1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r1), (forall (s1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_natrl s1), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r1 s1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_ntofrl r1 n) (l_e_st_eq_landau_n_rt_rp_r_ntofrl s1 m)))))). exact (fun (r1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r1) => (fun (s1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_natrl s1) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r1 s1) => l_e_isinv l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) l_e_st_eq_landau_n_rt_rp_r_ivr2_t31 r1 (l_e_st_eq_landau_n_rt_rp_r_natimage r1 n) s1 (l_e_st_eq_landau_n_rt_rp_r_natimage s1 m) i))))). Time Defined. (* constant 4231 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrlint : (forall (r1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r1), (forall (s1:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_natrl s1), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_ntofrl r1 n) (l_e_st_eq_landau_n_rt_rp_r_ntofrl s1 m)), l_e_st_eq_landau_n_rt_rp_r_is r1 s1))))). exact (fun (r1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r1) => (fun (s1:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_natrl s1) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_ntofrl r1 n) (l_e_st_eq_landau_n_rt_rp_r_ntofrl s1 m)) => l_e_isinve l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) l_e_st_eq_landau_n_rt_rp_r_ivr2_t31 r1 (l_e_st_eq_landau_n_rt_rp_r_natimage r1 n) s1 (l_e_st_eq_landau_n_rt_rp_r_natimage s1 m) i))))). Time Defined. (* constant 4232 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrlnt1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_ists1 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x) l_e_st_eq_landau_n_rt_rp_r_ivr2_t31 r (l_e_st_eq_landau_n_rt_rp_r_natimage r n))). Time Defined. (* constant 4233 *) Definition l_e_st_eq_landau_n_rt_rp_r_isrlnt2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n)) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n)) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 r n))). Time Defined. (* constant 4234 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_xn : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_soft l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt u) l_e_st_eq_landau_n_rt_rp_r_ivr2_t31 (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_imagei l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt u) x)). Time Defined. (* constant 4235 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t39 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_ivr2_xn x) (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_natrli x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_isinv l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt u) l_e_st_eq_landau_n_rt_rp_r_ivr2_t31 (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_imagei l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt u) x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_natimage (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_natrli x)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_rlofnt x))). Time Defined. (* constant 4236 *) Definition l_e_st_eq_landau_n_rt_rp_r_isntrl1 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is x (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_natrli x))). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_rt_rp_r_ivr2_xn x) (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_natrli x)) (l_e_isst1 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt u) l_e_st_eq_landau_n_rt_rp_r_ivr2_t31 x) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t39 x)). Time Defined. (* constant 4237 *) Definition l_e_st_eq_landau_n_rt_rp_r_isntrl2 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_natrli x)) x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat x (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_natrli x)) (l_e_st_eq_landau_n_rt_rp_r_isntrl1 x)). Time Defined. (* constant 4238 *) Definition l_e_st_eq_landau_n_rt_rp_r_intrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd x))). Time Defined. (* constant 4239 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t40 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a0), l_and (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a0) => l_andi (l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a0) a0ir i)))). Time Defined. (* constant 4240 *) Definition l_e_st_eq_landau_n_rt_rp_r_intrlin : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_intd a0), l_e_st_eq_landau_n_rt_rp_r_intrl r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_intd a0) => l_somei l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd x)) a0 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t40 r a0 a0ir i))))). Time Defined. (* constant 4241 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t41 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a)), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a) b)))))). Time Defined. (* constant 4242 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t42 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a)), l_e_st_eq_landau_n_rt_rp_intd a)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a) b)))))). Time Defined. (* constant 4243 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t43 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a)), l_e_st_eq_landau_n_rt_rp_intd a0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_and (l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd a)) => l_e_st_eq_landau_n_rt_rp_eqintd a a0 (l_e_st_eq_landau_n_rt_rp_r_isex r r a a0 (l_e_st_eq_landau_n_rt_rp_r_ivr2_t41 r a0 a0ir i a b) a0ir (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r)) (l_e_st_eq_landau_n_rt_rp_r_ivr2_t42 r a0 a0ir i a b))))))). Time Defined. (* constant 4244 *) Definition l_e_st_eq_landau_n_rt_rp_r_intrlex : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_intd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd x)) i (l_e_st_eq_landau_n_rt_rp_intd a0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) (l_e_st_eq_landau_n_rt_rp_intd x)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t43 r a0 a0ir i x t)))))). Time Defined. (* constant 4245 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t44 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_intd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_natintd a0 (l_e_st_eq_landau_n_rt_rp_r_natrlex r a0 a0ir n))))). Time Defined. (* constant 4246 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t45 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_r_intrl r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_r_intrlin r a0 a0ir (l_e_st_eq_landau_n_rt_rp_r_ivr2_t44 r a0 a0ir n))))). Time Defined. (* constant 4247 *) Definition l_e_st_eq_landau_n_rt_rp_r_natintrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_r_intrl r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_intrl r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t45 r x t n)))). Time Defined. (* constant 4248 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t46 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_natd a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_posintnatd a0 (l_e_st_eq_landau_n_rt_rp_r_posex r a0 a0ir p) (l_e_st_eq_landau_n_rt_rp_r_intrlex r a0 a0ir i)))))). Time Defined. (* constant 4249 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t47 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_r_natrl r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_r_natrlin r a0 a0ir (l_e_st_eq_landau_n_rt_rp_r_ivr2_t46 r a0 a0ir p i)))))). Time Defined. (* constant 4250 *) Definition l_e_st_eq_landau_n_rt_rp_r_posintnatrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_r_natrl r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_natrl r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t47 r x t p i))))). Time Defined. (* constant 4251 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t48 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i2:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_intd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i2:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_intdi0 a0 (l_e_st_eq_landau_n_rt_rp_r_0ex r a0 a0ir i2))))). Time Defined. (* constant 4252 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr2_t49 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i2:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_intrl r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i2:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_intrlin r a0 a0ir (l_e_st_eq_landau_n_rt_rp_r_ivr2_t48 r a0 a0ir i2))))). Time Defined. (* constant 4253 *) Definition l_e_st_eq_landau_n_rt_rp_r_intrli0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_intrl r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_intrl r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr2_t49 r x t i)))). Time Defined. (* constant 4254 *) Definition l_e_st_eq_landau_n_rt_rp_r_remark6 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r0), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_pofrp r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r0) => l_e_st_eq_landau_n_rt_rp_r_intrlin (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_remark6 r0 n))). Time Defined. (* constant 4255 *) Definition l_e_st_eq_landau_n_rt_rp_r_remark7 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (n:l_e_st_eq_landau_n_rt_rp_natrp r0), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_nofrp r0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (n:l_e_st_eq_landau_n_rt_rp_natrp r0) => l_e_st_eq_landau_n_rt_rp_r_intrlin (l_e_st_eq_landau_n_rt_rp_r_nofrp r0) (l_e_st_eq_landau_n_rt_rp_ndofrp r0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_ndofrp r0)) (l_e_st_eq_landau_n_rt_rp_remark7 r0 n))). Time Defined. (* constant 4256 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r174_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_ratd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_satzd174 a0 (l_e_st_eq_landau_n_rt_rp_r_intrlex r a0 a0ir i))))). Time Defined. (* constant 4257 *) Definition l_e_st_eq_landau_n_rt_rp_r_2r174_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_r_ratrl r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_r_ratrlin r a0 a0ir (l_e_st_eq_landau_n_rt_rp_r_2r174_t1 r a0 a0ir i))))). Time Defined. (* constant 4258 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz174 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_r_ratrl r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_ratrl r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_2r174_t2 r x t i)))). Time Defined. (* constant 4259 *) Definition l_e_st_eq_landau_n_rt_rp_r_plusdr : (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_pd x y))). Time Defined. (* constant 4260 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_plusdr a c) (l_e_st_eq_landau_n_rt_rp_r_plusdr b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_pd a c)) (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_pd b d)) (l_e_st_eq_landau_n_rt_rp_pd a c) (l_e_st_eq_landau_n_rt_rp_pd b d) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pd a c)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pd b d)) (l_e_st_eq_landau_n_rt_rp_eqpd12 a b c d e f))))))). Time Defined. (* constant 4261 *) Definition l_e_st_eq_landau_n_rt_rp_r_fplusdr : l_e_st_eq_landau_n_rt_rp_r_fixf2 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_plusdr. exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (v:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_eq x y) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_eq z v) => l_e_st_eq_landau_n_rt_rp_r_ivr3_t1 x y z v t u)))))). Time Defined. (* constant 4262 *) Definition l_e_st_eq_landau_n_rt_rp_r_pl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_indreal2 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_plusdr l_e_st_eq_landau_n_rt_rp_r_fplusdr r s)). Time Defined. (* constant 4263 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_pd a1 b1)) (l_e_st_eq_landau_n_rt_rp_r_pl r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isindreal2 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_plusdr l_e_st_eq_landau_n_rt_rp_r_fplusdr r s a1 b1 a1ir b1is)))))). Time Defined. (* constant 4264 *) Definition l_e_st_eq_landau_n_rt_rp_r_picp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_pd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_class (l_e_st_eq_landau_n_rt_rp_r_pl r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_pd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_class x)) (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_pd a1 b1)) (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pd a1 b1)) (l_e_st_eq_landau_n_rt_rp_r_ivr3_t2 r s a1 b1 a1ir b1is))))))). Time Defined. (* constant 4265 *) Definition l_e_st_eq_landau_n_rt_rp_r_ispl1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl x t) r s i)))). Time Defined. (* constant 4266 *) Definition l_e_st_eq_landau_n_rt_rp_r_ispl2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl t x) r s i)))). Time Defined. (* constant 4267 *) Definition l_e_st_eq_landau_n_rt_rp_r_ispl12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is t u), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is t u) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_ispl1 r s t i) (l_e_st_eq_landau_n_rt_rp_r_ispl2 t u s j))))))). Time Defined. (* constant 4268 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r175_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a1 b1) (l_e_st_eq_landau_n_rt_rp_pd b1 a1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_satzd175 a1 b1)))))). Time Defined. (* constant 4269 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r175_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl s r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl s r) (l_e_st_eq_landau_n_rt_rp_pd a1 b1) (l_e_st_eq_landau_n_rt_rp_pd b1 a1) (l_e_st_eq_landau_n_rt_rp_r_picp r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_picp s r b1 a1 b1is a1ir) (l_e_st_eq_landau_n_rt_rp_r_3r175_t1 r s a1 b1 a1ir b1is))))))). Time Defined. (* constant 4270 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz175 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl s r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r175_t2 r s x y t u)))))). Time Defined. (* constant 4271 *) Definition l_e_st_eq_landau_n_rt_rp_r_compl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_satz175 r s)). Time Defined. (* constant 4272 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd a1 b1) b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_pd01 a1 b1 (l_e_st_eq_landau_n_rt_rp_r_0ex r a1 a1ir i)))))))). Time Defined. (* constant 4273 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r s) s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_pl r s) s (l_e_st_eq_landau_n_rt_rp_pd a1 b1) b1 (l_e_st_eq_landau_n_rt_rp_r_picp r s a1 b1 a1ir b1is) b1is (l_e_st_eq_landau_n_rt_rp_r_ivr3_t3 r s a1 b1 a1ir b1is i)))))))). Time Defined. (* constant 4274 *) Definition l_e_st_eq_landau_n_rt_rp_r_pl01 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r s) s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r s) s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_ivr3_t4 r s x y t u i))))))). Time Defined. (* constant 4275 *) Definition l_e_st_eq_landau_n_rt_rp_r_pl02 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r s) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl s r) r (l_e_st_eq_landau_n_rt_rp_r_compl r s) (l_e_st_eq_landau_n_rt_rp_r_pl01 s r i)))). Time Defined. (* constant 4276 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_pd a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_ppd a1 b1 (l_e_st_eq_landau_n_rt_rp_r_posex r a1 a1ir p) (l_e_st_eq_landau_n_rt_rp_r_posex s b1 b1is q))))))))). Time Defined. (* constant 4277 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl r s))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_pd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_picp r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_ivr3_t5 r s a1 b1 a1ir b1is p q))))))))). Time Defined. (* constant 4278 *) Definition l_e_st_eq_landau_n_rt_rp_r_pospl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_ivr3_t6 r s x y t u p q)))))))). Time Defined. (* constant 4279 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_pd a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_npd a1 b1 (l_e_st_eq_landau_n_rt_rp_r_negex r a1 a1ir n) (l_e_st_eq_landau_n_rt_rp_r_negex s b1 b1is o))))))))). Time Defined. (* constant 4280 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_pl r s))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_negin (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_pd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_picp r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_ivr3_t7 r s a1 b1 a1ir b1is n o))))))))). Time Defined. (* constant 4281 *) Definition l_e_st_eq_landau_n_rt_rp_r_negpl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_pl r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_ivr3_t8 r s x y t u n o)))))))). Time Defined. (* constant 4282 *) Definition l_e_st_eq_landau_n_rt_rp_r_m0dr : (forall (x:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_m0d x)). Time Defined. (* constant 4283 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t5a : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0dr a) (l_e_st_eq_landau_n_rt_rp_r_m0dr b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_m0d a) (l_e_st_eq_landau_n_rt_rp_m0d b) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_m0d a)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_m0d b)) (l_e_st_eq_landau_n_rt_rp_eqm0d a b e)))). Time Defined. (* constant 4284 *) Definition l_e_st_eq_landau_n_rt_rp_r_fm0dr : l_e_st_eq_landau_n_rt_rp_r_fixf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_m0dr. exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_eq x y) => l_e_st_eq_landau_n_rt_rp_r_ivr3_t5a x y t))). Time Defined. (* constant 4285 *) Definition l_e_st_eq_landau_n_rt_rp_r_m0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_indreal l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_m0dr l_e_st_eq_landau_n_rt_rp_r_fm0dr r). Time Defined. (* constant 4286 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t6a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_m0d a0)) (l_e_st_eq_landau_n_rt_rp_r_m0 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_isindreal l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_m0dr l_e_st_eq_landau_n_rt_rp_r_fm0dr r a0 a0ir))). Time Defined. (* constant 4287 *) Definition l_e_st_eq_landau_n_rt_rp_r_micm0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_class (l_e_st_eq_landau_n_rt_rp_r_m0 r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_class x)) (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_m0d a0)) (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_m0d a0)) (l_e_st_eq_landau_n_rt_rp_r_ivr3_t6a r a0 a0ir)))). Time Defined. (* constant 4288 *) Definition l_e_st_eq_landau_n_rt_rp_r_ism0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_m0 x) r s i))). Time Defined. (* constant 4289 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t7a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_m0d a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_absnd a0 (l_e_st_eq_landau_n_rt_rp_r_negex r a0 a0ir n))))). Time Defined. (* constant 4290 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t8a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_m0 r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_aica r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_ivr3_t7a r a0 a0ir n))))). Time Defined. (* constant 4291 *) Definition l_e_st_eq_landau_n_rt_rp_r_absn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_m0 r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr3_t8a r x t n)))). Time Defined. (* constant 4292 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd a0) a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => l_e_st_eq_landau_n_rt_rp_absnnd a0 (l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd a0) (l_e_st_eq_landau_n_rt_rp_r_neg r) nn (fun (t:l_e_st_eq_landau_n_rt_rp_negd a0) => l_e_st_eq_landau_n_rt_rp_r_negin r a0 a0ir t)))))). Time Defined. (* constant 4293 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr3_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_abs r) r (l_e_st_eq_landau_n_rt_rp_absd a0) a0 (l_e_st_eq_landau_n_rt_rp_r_aica r a0 a0ir) a0ir (l_e_st_eq_landau_n_rt_rp_r_ivr3_t9 r a0 a0ir nn))))). Time Defined. (* constant 4294 *) Definition l_e_st_eq_landau_n_rt_rp_r_absnn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_ivr3_t10 r x t nn)))). Time Defined. (* constant 4295 *) Definition l_e_st_eq_landau_n_rt_rp_r_absp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_absnn r (l_e_st_eq_landau_n_rt_rp_r_pnotn r p))). Time Defined. (* constant 4296 *) Definition l_e_st_eq_landau_n_rt_rp_r_abs0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs r) r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_absnn r (l_e_st_eq_landau_n_rt_rp_r_0notn r i)) i)). Time Defined. (* constant 4297 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_m0d a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_satzd176a a0 (l_e_st_eq_landau_n_rt_rp_r_posex r a0 a0ir p))))). Time Defined. (* constant 4298 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_negin (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_3r176_t1 r a0 a0ir p))))). Time Defined. (* constant 4299 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz176a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r176_t2 r x t p)))). Time Defined. (* constant 4300 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_m0d a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_satzd176b a0 (l_e_st_eq_landau_n_rt_rp_r_0ex r a0 a0ir i))))). Time Defined. (* constant 4301 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_0in (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_3r176_t3 r a0 a0ir i))))). Time Defined. (* constant 4302 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz176b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r176_t4 r x t i)))). Time Defined. (* constant 4303 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_m0d a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_satzd176c a0 (l_e_st_eq_landau_n_rt_rp_r_negex r a0 a0ir n))))). Time Defined. (* constant 4304 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_3r176_t5 r a0 a0ir n))))). Time Defined. (* constant 4305 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz176c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r176_t6 r x t n)))). Time Defined. (* constant 4306 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_posd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_satzd176d a0 (l_e_st_eq_landau_n_rt_rp_r_negex (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir) n))))). Time Defined. (* constant 4307 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_pos r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_posin r a0 a0ir (l_e_st_eq_landau_n_rt_rp_r_3r176_t7 r a0 a0ir n))))). Time Defined. (* constant 4308 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz176d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_pos r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_pos r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r176_t8 r x t n)))). Time Defined. (* constant 4309 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_zero a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_satzd176e a0 (l_e_st_eq_landau_n_rt_rp_r_0ex (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir) i))))). Time Defined. (* constant 4310 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_0in r a0 a0ir (l_e_st_eq_landau_n_rt_rp_r_3r176_t9 r a0 a0ir i))))). Time Defined. (* constant 4311 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz176e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r176_t10 r x t i)))). Time Defined. (* constant 4312 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_negd a0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_satzd176f a0 (l_e_st_eq_landau_n_rt_rp_r_posex (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir) p))))). Time Defined. (* constant 4313 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r176_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_neg r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_negin r a0 a0ir (l_e_st_eq_landau_n_rt_rp_r_3r176_t11 r a0 a0ir p))))). Time Defined. (* constant 4314 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz176f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_neg r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_neg r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r176_t12 r x t p)))). Time Defined. (* constant 4315 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r177_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_m0d a0)) a0 (l_e_st_eq_landau_n_rt_rp_r_micm0 (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir)) a0ir (l_e_st_eq_landau_n_rt_rp_satzd177 a0)))). Time Defined. (* constant 4316 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz177 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r177_t1 r x t))). Time Defined. (* constant 4317 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz177a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r (l_e_st_eq_landau_n_rt_rp_r_satz177 r)). Time Defined. (* constant 4318 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz177b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s)) s (l_e_st_eq_landau_n_rt_rp_r_ism0 r (l_e_st_eq_landau_n_rt_rp_r_m0 s) i) (l_e_st_eq_landau_n_rt_rp_r_satz177 s)))). Time Defined. (* constant 4319 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz177c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 s)), l_e_st_eq_landau_n_rt_rp_r_is s (l_e_st_eq_landau_n_rt_rp_r_m0 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 r) s (l_e_st_eq_landau_n_rt_rp_r_satz177b r s i)))). Time Defined. (* constant 4320 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz177d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) s), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) => l_e_st_eq_landau_n_rt_rp_r_satz177c s r (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 r) s i)))). Time Defined. (* constant 4321 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz177e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 s) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_satz177d r s i)))). Time Defined. (* constant 4322 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r178_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_abs r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_m0d a0)) (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_r_aica (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir)) (l_e_st_eq_landau_n_rt_rp_r_aica r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_satzd178 a0)))). Time Defined. (* constant 4323 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz178 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_abs r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_abs r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r178_t1 r x t))). Time Defined. (* constant 4324 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz178a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_m0 r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_satz178 r)). Time Defined. (* constant 4325 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r179_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_m0 r)) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_0in (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_pd a0 (l_e_st_eq_landau_n_rt_rp_m0d a0)) (l_e_st_eq_landau_n_rt_rp_r_picp r (l_e_st_eq_landau_n_rt_rp_r_m0 r) a0 (l_e_st_eq_landau_n_rt_rp_m0d a0) a0ir (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir)) (l_e_st_eq_landau_n_rt_rp_satzd179 a0)))). Time Defined. (* constant 4326 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz179 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_m0 r)) l_e_st_eq_landau_n_rt_rp_r_0). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_m0 r)) l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r179_t1 r x t))). Time Defined. (* constant 4327 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz179a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) r) l_e_st_eq_landau_n_rt_rp_r_0). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) r) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_m0 r)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_m0 r) r) (l_e_st_eq_landau_n_rt_rp_r_satz179 r)). Time Defined. (* constant 4328 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r180_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_pd a1 b1)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_m0d a1) (l_e_st_eq_landau_n_rt_rp_m0d b1)) (l_e_st_eq_landau_n_rt_rp_r_micm0 (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_pd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_picp r s a1 b1 a1ir b1is)) (l_e_st_eq_landau_n_rt_rp_r_picp (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_m0d a1) (l_e_st_eq_landau_n_rt_rp_m0d b1) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_micm0 s b1 b1is)) (l_e_st_eq_landau_n_rt_rp_satzd180 a1 b1))))))). Time Defined. (* constant 4329 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz180 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r180_t1 r s x y t u)))))). Time Defined. (* constant 4330 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz180a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_satz180 r s))). Time Defined. (* constant 4331 *) Definition l_e_st_eq_landau_n_rt_rp_r_mn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_m0 s))). Time Defined. (* constant 4332 *) Definition l_e_st_eq_landau_n_rt_rp_r_micmn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_md a1 b1) (l_e_st_eq_landau_n_rt_rp_r_class (l_e_st_eq_landau_n_rt_rp_r_mn r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_picp r (l_e_st_eq_landau_n_rt_rp_r_m0 s) a1 (l_e_st_eq_landau_n_rt_rp_m0d b1) a1ir (l_e_st_eq_landau_n_rt_rp_r_micm0 s b1 b1is))))))). Time Defined. (* constant 4333 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismn1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r t) (l_e_st_eq_landau_n_rt_rp_r_mn s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_ispl1 r s (l_e_st_eq_landau_n_rt_rp_r_m0 t) i)))). Time Defined. (* constant 4334 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismn2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn t r) (l_e_st_eq_landau_n_rt_rp_r_mn t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s) t (l_e_st_eq_landau_n_rt_rp_r_ism0 r s i))))). Time Defined. (* constant 4335 *) Definition l_e_st_eq_landau_n_rt_rp_r_ismn12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is t u), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r t) (l_e_st_eq_landau_n_rt_rp_r_mn s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is t u) => l_e_st_eq_landau_n_rt_rp_r_ispl12 r s (l_e_st_eq_landau_n_rt_rp_r_m0 t) (l_e_st_eq_landau_n_rt_rp_r_m0 u) i (l_e_st_eq_landau_n_rt_rp_r_ism0 t u j))))))). Time Defined. (* constant 4336 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz181 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn r s)) (l_e_st_eq_landau_n_rt_rp_r_mn s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_mn s r) (l_e_st_eq_landau_n_rt_rp_r_satz180 r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s)) s (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_satz177 s)) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_m0 r) s))). Time Defined. (* constant 4337 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz181a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn s r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn s r)) (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_satz181 s r))). Time Defined. (* constant 4338 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)), l_e_st_eq_landau_n_rt_rp_mored a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)) => l_e_st_eq_landau_n_rt_rp_satzd182a a1 b1 (l_e_st_eq_landau_n_rt_rp_r_posex (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_md a1 b1) (l_e_st_eq_landau_n_rt_rp_r_micmn r s a1 b1 a1ir b1is) p)))))))). Time Defined. (* constant 4339 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)), l_e_st_eq_landau_n_rt_rp_r_more r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)) => l_e_st_eq_landau_n_rt_rp_r_morein r s a1 b1 a1ir b1is (l_e_st_eq_landau_n_rt_rp_r_3r182_t1 r s a1 b1 a1ir b1is p)))))))). Time Defined. (* constant 4340 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz182a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)), l_e_st_eq_landau_n_rt_rp_r_more r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_more r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r182_t2 r s x y t u p))))))). Time Defined. (* constant 4341 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_eq a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_satzd182b a1 b1 (l_e_st_eq_landau_n_rt_rp_r_0ex (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_md a1 b1) (l_e_st_eq_landau_n_rt_rp_r_micmn r s a1 b1 a1ir b1is) i)))))))). Time Defined. (* constant 4342 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_isin r s a1 b1 a1ir b1is (l_e_st_eq_landau_n_rt_rp_r_3r182_t3 r s a1 b1 a1ir b1is i)))))))). Time Defined. (* constant 4343 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz182b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r182_t4 r s x y t u i))))))). Time Defined. (* constant 4344 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)), l_e_st_eq_landau_n_rt_rp_lessd a1 b1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)) => l_e_st_eq_landau_n_rt_rp_satzd182c a1 b1 (l_e_st_eq_landau_n_rt_rp_r_negex (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_md a1 b1) (l_e_st_eq_landau_n_rt_rp_r_micmn r s a1 b1 a1ir b1is) n)))))))). Time Defined. (* constant 4345 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)), l_e_st_eq_landau_n_rt_rp_r_less r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)) => l_e_st_eq_landau_n_rt_rp_r_lessin r s a1 b1 a1ir b1is (l_e_st_eq_landau_n_rt_rp_r_3r182_t5 r s a1 b1 a1ir b1is n)))))))). Time Defined. (* constant 4346 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz182c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)), l_e_st_eq_landau_n_rt_rp_r_less r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_less r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r182_t6 r s x y t u n))))))). Time Defined. (* constant 4347 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_md a1 b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_satzd182d a1 b1 (l_e_st_eq_landau_n_rt_rp_r_moreex r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4348 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_md a1 b1) (l_e_st_eq_landau_n_rt_rp_r_micmn r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_3r182_t7 r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4349 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz182d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r182_t8 r s x y t u m))))))). Time Defined. (* constant 4350 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_md a1 b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_satzd182e a1 b1 (l_e_st_eq_landau_n_rt_rp_r_isex r s a1 b1 a1ir b1is i)))))))). Time Defined. (* constant 4351 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_0in (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_md a1 b1) (l_e_st_eq_landau_n_rt_rp_r_micmn r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_3r182_t9 r s a1 b1 a1ir b1is i)))))))). Time Defined. (* constant 4352 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz182e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r182_t10 r s x y t u i))))))). Time Defined. (* constant 4353 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_md a1 b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_satzd182f a1 b1 (l_e_st_eq_landau_n_rt_rp_r_lessex r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4354 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r182_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_negin (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_md a1 b1) (l_e_st_eq_landau_n_rt_rp_r_micmn r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_3r182_t11 r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4355 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz182f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r182_t12 r s x y t u l))))))). Time Defined. (* constant 4356 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r183_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_m0d a1) (l_e_st_eq_landau_n_rt_rp_m0d b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_satzd183a a1 b1 (l_e_st_eq_landau_n_rt_rp_r_moreex r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4357 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r183_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_lessin (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_m0d a1) (l_e_st_eq_landau_n_rt_rp_m0d b1) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_micm0 s b1 b1is) (l_e_st_eq_landau_n_rt_rp_r_3r183_t1 r s a1 b1 a1ir b1is m)))))))). Time Defined. (* constant 4358 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz183a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r183_t2 r s x y t u m))))))). Time Defined. (* constant 4359 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz183b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_ism0 r s i))). Time Defined. (* constant 4360 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r183_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_m0d a1) (l_e_st_eq_landau_n_rt_rp_m0d b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_satzd183c a1 b1 (l_e_st_eq_landau_n_rt_rp_r_lessex r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4361 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r183_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_morein (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_m0d a1) (l_e_st_eq_landau_n_rt_rp_m0d b1) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_micm0 s b1 b1is) (l_e_st_eq_landau_n_rt_rp_r_3r183_t3 r s a1 b1 a1ir b1is l)))))))). Time Defined. (* constant 4362 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz183c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_3r183_t4 r s x y t u l))))))). Time Defined. (* constant 4363 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz183d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)), l_e_st_eq_landau_n_rt_rp_r_more r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) => l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s)) s (l_e_st_eq_landau_n_rt_rp_r_satz177 r) (l_e_st_eq_landau_n_rt_rp_r_satz177 s) (l_e_st_eq_landau_n_rt_rp_r_satz183c (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s) l)))). Time Defined. (* constant 4364 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz183e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)), l_e_st_eq_landau_n_rt_rp_r_is r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s)) s (l_e_st_eq_landau_n_rt_rp_r_satz177a r) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s) i) (l_e_st_eq_landau_n_rt_rp_r_satz177 s)))). Time Defined. (* constant 4365 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz183f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)), l_e_st_eq_landau_n_rt_rp_r_less r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) => l_e_st_eq_landau_n_rt_rp_r_isless12 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s)) s (l_e_st_eq_landau_n_rt_rp_r_satz177 r) (l_e_st_eq_landau_n_rt_rp_r_satz177 s) (l_e_st_eq_landau_n_rt_rp_r_satz183a (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s) m)))). Time Defined. (* constant 4366 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_prop1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), Prop))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_and3 (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_pos t) (l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_mn s t))))). Time Defined. (* constant 4367 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_prop2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), Prop)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_3r184_prop1 r s x))). Time Defined. (* constant 4368 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_prop3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_3r184_prop2 r x)). Time Defined. (* constant 4369 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), Prop))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => l_and3 (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_eq a0 (l_e_st_eq_landau_n_rt_rp_md a b))))))). Time Defined. (* constant 4370 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a x))))). Time Defined. (* constant 4371 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_posd a))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_and3e1 (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_eq a0 (l_e_st_eq_landau_n_rt_rp_md a b)) p1))))))). Time Defined. (* constant 4372 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_posd b))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_and3e2 (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_eq a0 (l_e_st_eq_landau_n_rt_rp_md a b)) p1))))))). Time Defined. (* constant 4373 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_eq a0 (l_e_st_eq_landau_n_rt_rp_md a b)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_and3e3 (l_e_st_eq_landau_n_rt_rp_posd a) (l_e_st_eq_landau_n_rt_rp_posd b) (l_e_st_eq_landau_n_rt_rp_eq a0 (l_e_st_eq_landau_n_rt_rp_md a b)) p1))))))). Time Defined. (* constant 4374 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_ra : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), l_e_st_eq_landau_n_rt_rp_r_real))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => l_e_st_eq_landau_n_rt_rp_r_realof a))))). Time Defined. (* constant 4375 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_rb : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_r_real))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_e_st_eq_landau_n_rt_rp_r_realof b))))))). Time Defined. (* constant 4376 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_r_inn a (l_e_st_eq_landau_n_rt_rp_r_class (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_e_st_eq_landau_n_rt_rp_r_innclass a))))))). Time Defined. (* constant 4377 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_r_inn b (l_e_st_eq_landau_n_rt_rp_r_class (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_e_st_eq_landau_n_rt_rp_r_innclass b))))))). Time Defined. (* constant 4378 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2) (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_e_st_eq_landau_n_rt_rp_r_isin r (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2) (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1)) a0 (l_e_st_eq_landau_n_rt_rp_md a b) a0ir (l_e_st_eq_landau_n_rt_rp_r_micmn (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2) (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1) a b (l_e_st_eq_landau_n_rt_rp_r_3r184_t4 r a0 a0ir a p2 b p1) (l_e_st_eq_landau_n_rt_rp_r_3r184_t5 r a0 a0ir a p2 b p1)) (l_e_st_eq_landau_n_rt_rp_r_3r184_t3 r a0 a0ir a p2 b p1)))))))). Time Defined. (* constant 4379 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_r_3r184_prop1 r (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2) (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_and3i (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1)) (l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2) (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1))) (l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2) a (l_e_st_eq_landau_n_rt_rp_r_3r184_t4 r a0 a0ir a p2 b p1) (l_e_st_eq_landau_n_rt_rp_r_3r184_t1 r a0 a0ir a p2 b p1)) (l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1) b (l_e_st_eq_landau_n_rt_rp_r_3r184_t5 r a0 a0ir a p2 b p1) (l_e_st_eq_landau_n_rt_rp_r_3r184_t2 r a0 a0ir a p2 b p1)) (l_e_st_eq_landau_n_rt_rp_r_3r184_t6 r a0 a0ir a p2 b p1)))))))). Time Defined. (* constant 4380 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b), l_e_st_eq_landau_n_rt_rp_r_3r184_prop2 r (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a b) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_3r184_prop1 r (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2) x) (l_e_st_eq_landau_n_rt_rp_r_3r184_rb r a0 a0ir a p2 b p1) (l_e_st_eq_landau_n_rt_rp_r_3r184_t7 r a0 a0ir a p2 b p1)))))))). Time Defined. (* constant 4381 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), l_e_st_eq_landau_n_rt_rp_r_3r184_prop2 r (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a x) p2 (l_e_st_eq_landau_n_rt_rp_r_3r184_prop2 r (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_3r184_prop1d r a0 a0ir a x) => l_e_st_eq_landau_n_rt_rp_r_3r184_t8 r a0 a0ir a p2 x t))))))). Time Defined. (* constant 4382 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a), l_e_st_eq_landau_n_rt_rp_r_3r184_prop3 r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir a) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_3r184_prop2 r x) (l_e_st_eq_landau_n_rt_rp_r_3r184_ra r a0 a0ir a p2) (l_e_st_eq_landau_n_rt_rp_r_3r184_t9 r a0 a0ir a p2)))))). Time Defined. (* constant 4383 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r184_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_3r184_prop3 r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir x) (l_e_st_eq_landau_n_rt_rp_satzd184 a0) (l_e_st_eq_landau_n_rt_rp_r_3r184_prop3 r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_3r184_prop2d r a0 a0ir x) => l_e_st_eq_landau_n_rt_rp_r_3r184_t10 r a0 a0ir x t))))). Time Defined. (* constant 4384 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz184 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_some (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => l_and3 (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_pos y) (l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_mn x y))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_3r184_prop3 r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_3r184_t11 r x t))). Time Defined. (* constant 4385 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r185_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a3:l_e_st_eq_landau_n_rt_rp_dif), (forall (b3:l_e_st_eq_landau_n_rt_rp_dif), (forall (c3:l_e_st_eq_landau_n_rt_rp_dif), (forall (d3:l_e_st_eq_landau_n_rt_rp_dif), (forall (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md a3 b3) (l_e_st_eq_landau_n_rt_rp_md c3 d3)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a3 c3) (l_e_st_eq_landau_n_rt_rp_pd b3 d3)))))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)) => l_e_st_eq_landau_n_rt_rp_satzd185 a3 b3 c3 d3)))))))))))). Time Defined. (* constant 4386 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r185_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a3:l_e_st_eq_landau_n_rt_rp_dif), (forall (b3:l_e_st_eq_landau_n_rt_rp_dif), (forall (c3:l_e_st_eq_landau_n_rt_rp_dif), (forall (d3:l_e_st_eq_landau_n_rt_rp_dif), (forall (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_mn t u)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u)))))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_mn t u)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_md a3 b3) (l_e_st_eq_landau_n_rt_rp_md c3 d3)) (l_e_st_eq_landau_n_rt_rp_md (l_e_st_eq_landau_n_rt_rp_pd a3 c3) (l_e_st_eq_landau_n_rt_rp_pd b3 d3)) (l_e_st_eq_landau_n_rt_rp_r_picp (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_mn t u) (l_e_st_eq_landau_n_rt_rp_md a3 b3) (l_e_st_eq_landau_n_rt_rp_md c3 d3) (l_e_st_eq_landau_n_rt_rp_r_micmn r s a3 b3 a3ir b3is) (l_e_st_eq_landau_n_rt_rp_r_micmn t u c3 d3 c3it d3iu)) (l_e_st_eq_landau_n_rt_rp_r_micmn (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_pd a3 c3) (l_e_st_eq_landau_n_rt_rp_pd b3 d3) (l_e_st_eq_landau_n_rt_rp_r_picp r t a3 c3 a3ir c3it) (l_e_st_eq_landau_n_rt_rp_r_picp s u b3 d3 b3is d3iu)) (l_e_st_eq_landau_n_rt_rp_r_3r185_t1 r s t u a3 b3 c3 d3 a3ir b3is c3it d3iu))))))))))))). Time Defined. (* constant 4387 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz185 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_mn t u)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp4 r s t u (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_mn t u)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (v:l_e_st_eq_landau_n_rt_rp_dif) => (fun (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (vi:l_e_st_eq_landau_n_rt_rp_r_inn v (l_e_st_eq_landau_n_rt_rp_r_class u)) => l_e_st_eq_landau_n_rt_rp_r_3r185_t2 r s t u x y z v xi yi zi vi)))))))))))). Time Defined. (* constant 4388 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r186_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a2 b2) c2) (l_e_st_eq_landau_n_rt_rp_pd a2 (l_e_st_eq_landau_n_rt_rp_pd b2 c2))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_satzd186 a2 b2 c2))))))))). Time Defined. (* constant 4389 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r186_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl s t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pd a2 b2) c2) (l_e_st_eq_landau_n_rt_rp_pd a2 (l_e_st_eq_landau_n_rt_rp_pd b2 c2)) (l_e_st_eq_landau_n_rt_rp_r_picp (l_e_st_eq_landau_n_rt_rp_r_pl r s) t (l_e_st_eq_landau_n_rt_rp_pd a2 b2) c2 (l_e_st_eq_landau_n_rt_rp_r_picp r s a2 b2 a2ir b2is) c2it) (l_e_st_eq_landau_n_rt_rp_r_picp r (l_e_st_eq_landau_n_rt_rp_r_pl s t) a2 (l_e_st_eq_landau_n_rt_rp_pd b2 c2) a2ir (l_e_st_eq_landau_n_rt_rp_r_picp s t b2 c2 b2is c2it)) (l_e_st_eq_landau_n_rt_rp_r_3r186_t1 r s t a2 b2 c2 a2ir b2is c2it)))))))))). Time Defined. (* constant 4390 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz186 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl s t))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_3r186_t2 r s t x y z u v w))))))))). Time Defined. (* constant 4391 *) Definition l_e_st_eq_landau_n_rt_rp_r_asspl1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_satz186 r s t))). Time Defined. (* constant 4392 *) Definition l_e_st_eq_landau_n_rt_rp_r_asspl2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl r s) t)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_r_satz186 r s t)))). Time Defined. (* constant 4393 *) Definition l_e_st_eq_landau_n_rt_rp_r_plmn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn r s) s) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn r s) s) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 s) s)) r (l_e_st_eq_landau_n_rt_rp_r_asspl1 r (l_e_st_eq_landau_n_rt_rp_r_m0 s) s) (l_e_st_eq_landau_n_rt_rp_r_pl02 r (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 s) s) (l_e_st_eq_landau_n_rt_rp_r_satz179a s)))). Time Defined. (* constant 4394 *) Definition l_e_st_eq_landau_n_rt_rp_r_mnpl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl r s) s) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl r s) s) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_pl s (l_e_st_eq_landau_n_rt_rp_r_m0 s))) r (l_e_st_eq_landau_n_rt_rp_r_asspl1 r s (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_pl02 r (l_e_st_eq_landau_n_rt_rp_r_pl s (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_satz179 s)))). Time Defined. (* constant 4395 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz187a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s (l_e_st_eq_landau_n_rt_rp_r_mn r s)) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl s (l_e_st_eq_landau_n_rt_rp_r_mn r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn r s) s) r (l_e_st_eq_landau_n_rt_rp_r_compl s (l_e_st_eq_landau_n_rt_rp_r_mn r s)) (l_e_st_eq_landau_n_rt_rp_r_plmn r s))). Time Defined. (* constant 4396 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz187b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r) (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_satz187a r s))). Time Defined. (* constant 4397 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz187c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) x)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x s) s) x (l_e_st_eq_landau_n_rt_rp_r_ismn1 r (l_e_st_eq_landau_n_rt_rp_r_pl x s) s (l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_pl x s) (l_e_st_eq_landau_n_rt_rp_r_pl s x) i (l_e_st_eq_landau_n_rt_rp_r_compl s x))) (l_e_st_eq_landau_n_rt_rp_r_mnpl x s))))). Time Defined. (* constant 4398 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz187d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r), l_e_st_eq_landau_n_rt_rp_r_is x (l_e_st_eq_landau_n_rt_rp_r_mn r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn r s) x (l_e_st_eq_landau_n_rt_rp_r_satz187c r s x i))))). Time Defined. (* constant 4399 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz187e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl x s) r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn r s) x)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl x s) r) => l_e_st_eq_landau_n_rt_rp_r_satz187c r s x (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl s x) (l_e_st_eq_landau_n_rt_rp_r_pl x s) r (l_e_st_eq_landau_n_rt_rp_r_compl s x) i))))). Time Defined. (* constant 4400 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz187f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl x s) r), l_e_st_eq_landau_n_rt_rp_r_is x (l_e_st_eq_landau_n_rt_rp_r_mn r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl x s) r) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn r s) x (l_e_st_eq_landau_n_rt_rp_r_satz187e r s x i))))). Time Defined. (* constant 4401 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r187_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s y) r), l_e_st_eq_landau_n_rt_rp_r_is x y)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s y) r) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real x y (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_satz187c r s x i) (l_e_st_eq_landau_n_rt_rp_r_satz187c r s y j))))))). Time Defined. (* constant 4402 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r187_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_amone l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s y) r) => l_e_st_eq_landau_n_rt_rp_r_3r187_t1 r s x y t u)))))). Time Defined. (* constant 4403 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz187 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_one (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_onei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl s x) r) (l_e_st_eq_landau_n_rt_rp_r_3r187_t2 r s) (l_e_st_eq_landau_n_rt_rp_r_satz187b r s))). Time Defined. (* constant 4404 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_mored a2 b2)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_satzd188a a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_moreex (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_pd a2 c2) (l_e_st_eq_landau_n_rt_rp_pd b2 c2) (l_e_st_eq_landau_n_rt_rp_r_picp r t a2 c2 a2ir c2it) (l_e_st_eq_landau_n_rt_rp_r_picp s t b2 c2 b2is c2it) m))))))))))). Time Defined. (* constant 4405 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_r_more r s)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_r_morein r s a2 b2 a2ir b2is (l_e_st_eq_landau_n_rt_rp_r_3r188_t1 r s t a2 b2 c2 a2ir b2is c2it m))))))))))). Time Defined. (* constant 4406 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_r_more r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_more r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_3r188_t2 r s t x y z u v w m)))))))))). Time Defined. (* constant 4407 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_eq a2 b2)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_satzd188b a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_pd a2 c2) (l_e_st_eq_landau_n_rt_rp_pd b2 c2) (l_e_st_eq_landau_n_rt_rp_r_picp r t a2 c2 a2ir c2it) (l_e_st_eq_landau_n_rt_rp_r_picp s t b2 c2 b2is c2it) i))))))))))). Time Defined. (* constant 4408 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_r_is r s)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_r_isin r s a2 b2 a2ir b2is (l_e_st_eq_landau_n_rt_rp_r_3r188_t3 r s t a2 b2 c2 a2ir b2is c2it i))))))))))). Time Defined. (* constant 4409 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_r_is r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_is r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_3r188_t4 r s t x y z u v w i)))))))))). Time Defined. (* constant 4410 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_lessd a2 b2)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_satzd188c a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_lessex (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_pd a2 c2) (l_e_st_eq_landau_n_rt_rp_pd b2 c2) (l_e_st_eq_landau_n_rt_rp_r_picp r t a2 c2 a2ir c2it) (l_e_st_eq_landau_n_rt_rp_r_picp s t b2 c2 b2is c2it) l))))))))))). Time Defined. (* constant 4411 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_r_less r s)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_r_lessin r s a2 b2 a2ir b2is (l_e_st_eq_landau_n_rt_rp_r_3r188_t5 r s t a2 b2 c2 a2ir b2is c2it l))))))))))). Time Defined. (* constant 4412 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)), l_e_st_eq_landau_n_rt_rp_r_less r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_less r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_3r188_t6 r s t x y z u v w l)))))))))). Time Defined. (* constant 4413 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_pd a2 c2) (l_e_st_eq_landau_n_rt_rp_pd b2 c2))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_satzd188d a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_moreex r s a2 b2 a2ir b2is m))))))))))). Time Defined. (* constant 4414 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_morein (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_pd a2 c2) (l_e_st_eq_landau_n_rt_rp_pd b2 c2) (l_e_st_eq_landau_n_rt_rp_r_picp r t a2 c2 a2ir c2it) (l_e_st_eq_landau_n_rt_rp_r_picp s t b2 c2 b2is c2it) (l_e_st_eq_landau_n_rt_rp_r_3r188_t7 r s t a2 b2 c2 a2ir b2is c2it m))))))))))). Time Defined. (* constant 4415 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_3r188_t8 r s t x y z u v w m)))))))))). Time Defined. (* constant 4416 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_ispl1 r s t i)))). Time Defined. (* constant 4417 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_pd a2 c2) (l_e_st_eq_landau_n_rt_rp_pd b2 c2))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_satzd188f a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_lessex r s a2 b2 a2ir b2is l))))))))))). Time Defined. (* constant 4418 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r188_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_lessin (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_pd a2 c2) (l_e_st_eq_landau_n_rt_rp_pd b2 c2) (l_e_st_eq_landau_n_rt_rp_r_picp r t a2 c2 a2ir c2it) (l_e_st_eq_landau_n_rt_rp_r_picp s t b2 c2 b2is c2it) (l_e_st_eq_landau_n_rt_rp_r_3r188_t9 r s t a2 b2 c2 a2ir b2is c2it l))))))))))). Time Defined. (* constant 4419 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_3r188_t10 r s t x y z u v w l)))))))))). Time Defined. (* constant 4420 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188g : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s)), l_e_st_eq_landau_n_rt_rp_r_more r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s)) => l_e_st_eq_landau_n_rt_rp_r_satz188a r s t (l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl t s) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_r_compl t r) (l_e_st_eq_landau_n_rt_rp_r_compl t s) m))))). Time Defined. (* constant 4421 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188h : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s)), l_e_st_eq_landau_n_rt_rp_r_is r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s)) => l_e_st_eq_landau_n_rt_rp_r_satz188b r s t (l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_r_compl r t) i (l_e_st_eq_landau_n_rt_rp_r_compl t s)))))). Time Defined. (* constant 4422 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188j : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s)), l_e_st_eq_landau_n_rt_rp_r_less r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s)) => l_e_st_eq_landau_n_rt_rp_r_satz188c r s t (l_e_st_eq_landau_n_rt_rp_r_isless12 (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl t s) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_r_compl t r) (l_e_st_eq_landau_n_rt_rp_r_compl t s) l))))). Time Defined. (* constant 4423 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188k : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_r_pl t s) (l_e_st_eq_landau_n_rt_rp_r_compl r t) (l_e_st_eq_landau_n_rt_rp_r_compl s t) (l_e_st_eq_landau_n_rt_rp_r_satz188d r s t m))))). Time Defined. (* constant 4424 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188l : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_ispl2 r s t i)))). Time Defined. (* constant 4425 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188m : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_isless12 (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_r_pl t s) (l_e_st_eq_landau_n_rt_rp_r_compl r t) (l_e_st_eq_landau_n_rt_rp_r_compl s t) (l_e_st_eq_landau_n_rt_rp_r_satz188f r s t l))))). Time Defined. (* constant 4426 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188n : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more t u), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more t u) => l_e_st_eq_landau_n_rt_rp_r_ismore2 (l_e_st_eq_landau_n_rt_rp_r_pl r u) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_ispl1 r s u i) (l_e_st_eq_landau_n_rt_rp_r_satz188k t u r m))))))). Time Defined. (* constant 4427 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188o : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more t u), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl u s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more t u) => l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_pl u s) (l_e_st_eq_landau_n_rt_rp_r_compl r t) (l_e_st_eq_landau_n_rt_rp_r_compl s u) (l_e_st_eq_landau_n_rt_rp_r_satz188n r s t u i m))))))). Time Defined. (* constant 4428 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188p : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less t u), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less t u) => l_e_st_eq_landau_n_rt_rp_r_isless2 (l_e_st_eq_landau_n_rt_rp_r_pl r u) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_ispl1 r s u i) (l_e_st_eq_landau_n_rt_rp_r_satz188m t u r l))))))). Time Defined. (* constant 4429 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz188q : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less t u), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl u s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less t u) => l_e_st_eq_landau_n_rt_rp_r_isless12 (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_pl u s) (l_e_st_eq_landau_n_rt_rp_r_compl r t) (l_e_st_eq_landau_n_rt_rp_r_compl s u) (l_e_st_eq_landau_n_rt_rp_r_satz188p r s t u i l))))))). Time Defined. (* constant 4430 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz189 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_more t u), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_more t u) => l_e_st_eq_landau_n_rt_rp_r_trmore (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s t) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_satz188d r s t m) (l_e_st_eq_landau_n_rt_rp_r_satz188k t u s n))))))). Time Defined. (* constant 4431 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz189a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less t u), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less t u) => l_e_st_eq_landau_n_rt_rp_r_lemma1 (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_satz189 s r u t (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l) (l_e_st_eq_landau_n_rt_rp_r_lemma2 t u k)))))))). Time Defined. (* constant 4432 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz190a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_more t u), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_more t u) => l_orapp (l_e_st_eq_landau_n_rt_rp_r_more r s) (l_e_st_eq_landau_n_rt_rp_r_is r s) (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u)) m (fun (v:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_satz189 r s t u v n) (fun (v:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_satz188n r s t u v n))))))). Time Defined. (* constant 4433 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz190b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis t u), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis t u) => l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_pl t r) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl u s) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_compl t r) (l_e_st_eq_landau_n_rt_rp_r_compl u s) (l_e_st_eq_landau_n_rt_rp_r_satz190a t u r s n m))))))). Time Defined. (* constant 4434 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz190c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less t u), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less t u) => l_e_st_eq_landau_n_rt_rp_r_lemma1 (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_satz190a s r u t (l_e_st_eq_landau_n_rt_rp_r_satz168b r s l) (l_e_st_eq_landau_n_rt_rp_r_lemma2 t u k)))))))). Time Defined. (* constant 4435 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz190d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis t u), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis t u) => l_e_st_eq_landau_n_rt_rp_r_lemma1 (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_satz190b s r u t (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l) (l_e_st_eq_landau_n_rt_rp_r_satz168b t u k)))))))). Time Defined. (* constant 4436 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r191_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a3:l_e_st_eq_landau_n_rt_rp_dif), (forall (b3:l_e_st_eq_landau_n_rt_rp_dif), (forall (c3:l_e_st_eq_landau_n_rt_rp_dif), (forall (d3:l_e_st_eq_landau_n_rt_rp_dif), (forall (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis t u), l_e_st_eq_landau_n_rt_rp_moreq (l_e_st_eq_landau_n_rt_rp_pd a3 c3) (l_e_st_eq_landau_n_rt_rp_pd b3 d3))))))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis t u) => l_e_st_eq_landau_n_rt_rp_satzd191 a3 b3 c3 d3 (l_e_st_eq_landau_n_rt_rp_r_moreisex r s a3 b3 a3ir b3is m) (l_e_st_eq_landau_n_rt_rp_r_moreisex t u c3 d3 c3it d3iu n))))))))))))))). Time Defined. (* constant 4437 *) Definition l_e_st_eq_landau_n_rt_rp_r_3r191_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a3:l_e_st_eq_landau_n_rt_rp_dif), (forall (b3:l_e_st_eq_landau_n_rt_rp_dif), (forall (c3:l_e_st_eq_landau_n_rt_rp_dif), (forall (d3:l_e_st_eq_landau_n_rt_rp_dif), (forall (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis t u), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis t u) => l_e_st_eq_landau_n_rt_rp_r_moreisin (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_pd a3 c3) (l_e_st_eq_landau_n_rt_rp_pd b3 d3) (l_e_st_eq_landau_n_rt_rp_r_picp r t a3 c3 a3ir c3it) (l_e_st_eq_landau_n_rt_rp_r_picp s u b3 d3 b3is d3iu) (l_e_st_eq_landau_n_rt_rp_r_3r191_t1 r s t u a3 b3 c3 d3 a3ir b3is c3it d3iu m n))))))))))))))). Time Defined. (* constant 4438 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz191 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis t u), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis t u) => l_e_st_eq_landau_n_rt_rp_r_realapp4 r s t u (l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (v:l_e_st_eq_landau_n_rt_rp_dif) => (fun (xi:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (yi:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (vi:l_e_st_eq_landau_n_rt_rp_r_inn v (l_e_st_eq_landau_n_rt_rp_r_class u)) => l_e_st_eq_landau_n_rt_rp_r_3r191_t2 r s t u x y z v xi yi zi vi m n)))))))))))))). Time Defined. (* constant 4439 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz191a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s), (forall (k:l_e_st_eq_landau_n_rt_rp_r_lessis t u), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_pl s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis r s) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_lessis t u) => l_e_st_eq_landau_n_rt_rp_r_satz168a (l_e_st_eq_landau_n_rt_rp_r_pl s u) (l_e_st_eq_landau_n_rt_rp_r_pl r t) (l_e_st_eq_landau_n_rt_rp_r_satz191 s r u t (l_e_st_eq_landau_n_rt_rp_r_satz168b r s l) (l_e_st_eq_landau_n_rt_rp_r_satz168b t u k)))))))). Time Defined. (* constant 4440 *) Definition l_e_st_eq_landau_n_rt_rp_r_timesdr : (forall (x:l_e_st_eq_landau_n_rt_rp_dif), (forall (y:l_e_st_eq_landau_n_rt_rp_dif), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_td x y))). Time Defined. (* constant 4441 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr4_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (b:l_e_st_eq_landau_n_rt_rp_dif), (forall (c:l_e_st_eq_landau_n_rt_rp_dif), (forall (d:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq a b), (forall (f:l_e_st_eq_landau_n_rt_rp_eq c d), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_timesdr a c) (l_e_st_eq_landau_n_rt_rp_r_timesdr b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq a b) => (fun (f:l_e_st_eq_landau_n_rt_rp_eq c d) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_td b d)) (l_e_st_eq_landau_n_rt_rp_td a c) (l_e_st_eq_landau_n_rt_rp_td b d) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_td a c)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_td b d)) (l_e_st_eq_landau_n_rt_rp_eqtd12 a b c d e f))))))). Time Defined. (* constant 4442 *) Definition l_e_st_eq_landau_n_rt_rp_r_ftimesdr : l_e_st_eq_landau_n_rt_rp_r_fixf2 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_timesdr. exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (v:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_eq x y) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_eq z v) => l_e_st_eq_landau_n_rt_rp_r_ivr4_t1 x y z v t u)))))). Time Defined. (* constant 4443 *) Definition l_e_st_eq_landau_n_rt_rp_r_ts : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_indreal2 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_timesdr l_e_st_eq_landau_n_rt_rp_r_ftimesdr r s)). Time Defined. (* constant 4444 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr4_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_td a1 b1)) (l_e_st_eq_landau_n_rt_rp_r_ts r s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isindreal2 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_timesdr l_e_st_eq_landau_n_rt_rp_r_ftimesdr r s a1 b1 a1ir b1is)))))). Time Defined. (* constant 4445 *) Definition l_e_st_eq_landau_n_rt_rp_r_tict : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_class (l_e_st_eq_landau_n_rt_rp_r_ts r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_inn (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_class x)) (l_e_st_eq_landau_n_rt_rp_r_realof (l_e_st_eq_landau_n_rt_rp_td a1 b1)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_td a1 b1)) (l_e_st_eq_landau_n_rt_rp_r_ivr4_t2 r s a1 b1 a1ir b1is))))))). Time Defined. (* constant 4446 *) Definition l_e_st_eq_landau_n_rt_rp_r_ists1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ts x t) r s i)))). Time Defined. (* constant 4447 *) Definition l_e_st_eq_landau_n_rt_rp_r_ists2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ts t x) r s i)))). Time Defined. (* constant 4448 *) Definition l_e_st_eq_landau_n_rt_rp_r_ists12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is t u), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s u))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is t u) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t) (l_e_st_eq_landau_n_rt_rp_r_ts s u) (l_e_st_eq_landau_n_rt_rp_r_ists1 r s t i) (l_e_st_eq_landau_n_rt_rp_r_ists2 t u s j))))))). Time Defined. (* constant 4449 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r192_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a1 b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_satzd192a a1 b1 (l_e_st_eq_landau_n_rt_rp_r_0ex r a1 a1ir i)))))))). Time Defined. (* constant 4450 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r192_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_0in (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_4r192_t1 r s a1 b1 a1ir b1is i)))))))). Time Defined. (* constant 4451 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz192a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r192_t2 r s x y t u i))))))). Time Defined. (* constant 4452 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r192_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_zero (l_e_st_eq_landau_n_rt_rp_td a1 b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_satzd192b a1 b1 (l_e_st_eq_landau_n_rt_rp_r_0ex s b1 b1is i)))))))). Time Defined. (* constant 4453 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r192_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_0in (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_4r192_t3 r s a1 b1 a1ir b1is i)))))))). Time Defined. (* constant 4454 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz192b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r192_t4 r s x y t u i))))))). Time Defined. (* constant 4455 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r192_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0), l_or (l_e_st_eq_landau_n_rt_rp_zero a1) (l_e_st_eq_landau_n_rt_rp_zero b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_satzd192c a1 b1 (l_e_st_eq_landau_n_rt_rp_r_0ex (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) i)))))))). Time Defined. (* constant 4456 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r192_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0), l_or (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_zero a1) (l_e_st_eq_landau_n_rt_rp_zero b1) (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_4r192_t5 r s a1 b1 a1ir b1is i) (fun (t:l_e_st_eq_landau_n_rt_rp_zero a1) => l_e_st_eq_landau_n_rt_rp_r_0in r a1 a1ir t) (fun (t:l_e_st_eq_landau_n_rt_rp_zero b1) => l_e_st_eq_landau_n_rt_rp_r_0in s b1 b1is t)))))))). Time Defined. (* constant 4457 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz192c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0), l_or (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_or (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r192_t6 r s x y t u i))))))). Time Defined. (* constant 4458 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz192d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0), (forall (o:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_or (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) (l_or_th3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) n o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz192c r s t))))). Time Defined. (* constant 4459 *) Definition l_e_st_eq_landau_n_rt_rp_r_ts01 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz192a r s i))). Time Defined. (* constant 4460 *) Definition l_e_st_eq_landau_n_rt_rp_r_ts02 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz192b r s i))). Time Defined. (* constant 4461 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r193_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a1 b1)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_satzd193 a1 b1)))))). Time Defined. (* constant 4462 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r193_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_absd (l_e_st_eq_landau_n_rt_rp_td a1 b1)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1)) (l_e_st_eq_landau_n_rt_rp_r_aica (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is)) (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1) (l_e_st_eq_landau_n_rt_rp_r_aica r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_aica s b1 b1is)) (l_e_st_eq_landau_n_rt_rp_r_4r193_t1 r s a1 b1 a1ir b1is))))))). Time Defined. (* constant 4463 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz193 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r193_t2 r s x y t u)))))). Time Defined. (* constant 4464 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz193a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_ts r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_r_satz193 r s))). Time Defined. (* constant 4465 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r194_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_td b1 a1))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_satzd194 a1 b1)))))). Time Defined. (* constant 4466 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r194_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts s r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts s r) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_td b1 a1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_tict s r b1 a1 b1is a1ir) (l_e_st_eq_landau_n_rt_rp_r_4r194_t1 r s a1 b1 a1ir b1is))))))). Time Defined. (* constant 4467 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz194 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts s r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r194_t2 r s x y t u)))))). Time Defined. (* constant 4468 *) Definition l_e_st_eq_landau_n_rt_rp_r_comts : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts s r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_satz194 r s)). Time Defined. (* constant 4469 *) Definition l_e_st_eq_landau_n_rt_rp_r_1rl : l_e_st_eq_landau_n_rt_rp_r_real. exact (l_e_st_eq_landau_n_rt_rp_r_realof l_e_st_eq_landau_n_rt_rp_1df). Time Defined. (* constant 4470 *) Definition l_e_st_eq_landau_n_rt_rp_r_pos1 : l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_1rl. exact (l_e_st_eq_landau_n_rt_rp_r_posin l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_1df (l_e_st_eq_landau_n_rt_rp_r_innclass l_e_st_eq_landau_n_rt_rp_1df) (l_e_st_eq_landau_n_rt_rp_posdirp l_e_st_eq_landau_n_rt_rp_1rp)). Time Defined. (* constant 4471 *) Definition l_e_st_eq_landau_n_rt_rp_r_natrl1 : l_e_st_eq_landau_n_rt_rp_r_natrl l_e_st_eq_landau_n_rt_rp_r_1rl. exact (l_e_st_eq_landau_n_rt_rp_r_natrli l_e_st_eq_landau_n_1). Time Defined. (* constant 4472 *) Definition l_e_st_eq_landau_n_rt_rp_r_intrl1 : l_e_st_eq_landau_n_rt_rp_r_intrl l_e_st_eq_landau_n_rt_rp_r_1rl. exact (l_e_st_eq_landau_n_rt_rp_r_natintrl l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_natrl1). Time Defined. (* constant 4473 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r195_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 l_e_st_eq_landau_n_rt_rp_1df) a0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_satzd195 a0))). Time Defined. (* constant 4474 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r195_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_1rl) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_1rl) r (l_e_st_eq_landau_n_rt_rp_td a0 l_e_st_eq_landau_n_rt_rp_1df) a0 (l_e_st_eq_landau_n_rt_rp_r_tict r l_e_st_eq_landau_n_rt_rp_r_1rl a0 l_e_st_eq_landau_n_rt_rp_1df a0ir (l_e_st_eq_landau_n_rt_rp_r_innclass l_e_st_eq_landau_n_rt_rp_1df)) a0ir (l_e_st_eq_landau_n_rt_rp_r_4r195_t1 r a0 a0ir)))). Time Defined. (* constant 4475 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz195 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_1rl) r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_1rl) r) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_4r195_t2 r x t))). Time Defined. (* constant 4476 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz195a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_1rl)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_1rl) r (l_e_st_eq_landau_n_rt_rp_r_satz195 r)). Time Defined. (* constant 4477 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz195b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r) r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_1rl) r (l_e_st_eq_landau_n_rt_rp_r_comts l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_satz195 r)). Time Defined. (* constant 4478 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz195c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r) r (l_e_st_eq_landau_n_rt_rp_r_satz195b r)). Time Defined. (* constant 4479 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz196a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ists12 (l_e_st_eq_landau_n_rt_rp_r_abs r) r (l_e_st_eq_landau_n_rt_rp_r_abs s) s (l_e_st_eq_landau_n_rt_rp_r_absp r p) (l_e_st_eq_landau_n_rt_rp_r_absp s q)))))). Time Defined. (* constant 4480 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_satzd196b a1 b1 (l_e_st_eq_landau_n_rt_rp_r_negex r a1 a1ir n) (l_e_st_eq_landau_n_rt_rp_r_negex s b1 b1is o))))))))). Time Defined. (* constant 4481 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1)) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1) (l_e_st_eq_landau_n_rt_rp_r_aica r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_aica s b1 b1is)) (l_e_st_eq_landau_n_rt_rp_r_4r196_t1 r s a1 b1 a1ir b1is n o))))))))). Time Defined. (* constant 4482 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz196b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r196_t2 r s x y t u n o)))))))). Time Defined. (* constant 4483 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t1a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_satzd196c a1 b1 (l_e_st_eq_landau_n_rt_rp_r_posex r a1 a1ir p) (l_e_st_eq_landau_n_rt_rp_r_negex s b1 b1is n))))))))). Time Defined. (* constant 4484 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t2a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1))) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_micm0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1)) (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1) (l_e_st_eq_landau_n_rt_rp_r_aica r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_aica s b1 b1is))) (l_e_st_eq_landau_n_rt_rp_r_4r196_t1a r s a1 b1 a1ir b1is p n))))))))). Time Defined. (* constant 4485 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz196c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r196_t2a r s x y t u p n)))))))). Time Defined. (* constant 4486 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz196d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts s r) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_r_abs r))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) (l_e_st_eq_landau_n_rt_rp_r_comts r s) (l_e_st_eq_landau_n_rt_rp_r_satz196c s r p n) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_r_abs r)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_r_abs r))))))). Time Defined. (* constant 4487 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), l_not (l_e_st_eq_landau_n_rt_rp_zero a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_zero a0) (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) n (fun (t:l_e_st_eq_landau_n_rt_rp_zero a0) => l_e_st_eq_landau_n_rt_rp_r_0in r a0 a0ir t))))). Time Defined. (* constant 4488 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))), l_or (l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)) (l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_negd b1))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) => l_e_st_eq_landau_n_rt_rp_satzd196e a1 b1 (l_e_st_eq_landau_n_rt_rp_r_4r196_t3 r a1 a1ir n) (l_e_st_eq_landau_n_rt_rp_r_4r196_t3 s b1 b1is o) (l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1)) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1) (l_e_st_eq_landau_n_rt_rp_r_aica r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_aica s b1 b1is)) i)))))))))). Time Defined. (* constant 4489 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)), l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_pos s))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_posin r a1 a1ir (l_ande1 (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_posd b1) a)) (l_e_st_eq_landau_n_rt_rp_r_posin s b1 b1is (l_ande2 (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_posd b1) a)))))))))))). Time Defined. (* constant 4490 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)), l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_neg s))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_negin r a1 a1ir (l_ande1 (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_negd b1) a)) (l_e_st_eq_landau_n_rt_rp_r_negin s b1 b1is (l_ande2 (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_negd b1) a)))))))))))). Time Defined. (* constant 4491 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))), l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_pos s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_neg s))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) => l_or_th9 (l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)) (l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_pos s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_4r196_t4 r s a1 b1 a1ir b1is n o i) (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)) => l_e_st_eq_landau_n_rt_rp_r_4r196_t5 r s a1 b1 a1ir b1is n o i t) (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)) => l_e_st_eq_landau_n_rt_rp_r_4r196_t6 r s a1 b1 a1ir b1is n o i t)))))))))). Time Defined. (* constant 4492 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz196e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))), l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_pos s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_neg s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_pos s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_neg s))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r196_t7 r s x y t u n o i))))))))). Time Defined. (* constant 4493 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))), l_or (l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)) (l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_posd b1))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))) => l_e_st_eq_landau_n_rt_rp_satzd196f a1 b1 (l_e_st_eq_landau_n_rt_rp_r_4r196_t3 r a1 a1ir n) (l_e_st_eq_landau_n_rt_rp_r_4r196_t3 s b1 b1is o) (l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1))) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_micm0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1)) (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_absd a1) (l_e_st_eq_landau_n_rt_rp_absd b1) (l_e_st_eq_landau_n_rt_rp_r_aica r a1 a1ir) (l_e_st_eq_landau_n_rt_rp_r_aica s b1 b1is))) i)))))))))). Time Defined. (* constant 4494 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)), l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg s))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_posin r a1 a1ir (l_ande1 (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_negd b1) a)) (l_e_st_eq_landau_n_rt_rp_r_negin s b1 b1is (l_ande2 (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_negd b1) a)))))))))))). Time Defined. (* constant 4495 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)), l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_pos s))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_negin r a1 a1ir (l_ande1 (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_posd b1) a)) (l_e_st_eq_landau_n_rt_rp_r_posin s b1 b1is (l_ande2 (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_posd b1) a)))))))))))). Time Defined. (* constant 4496 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))), l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_pos s))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))) => l_or_th9 (l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)) (l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_pos s)) (l_e_st_eq_landau_n_rt_rp_r_4r196_t8 r s a1 b1 a1ir b1is n o i) (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_posd a1) (l_e_st_eq_landau_n_rt_rp_negd b1)) => l_e_st_eq_landau_n_rt_rp_r_4r196_t9 r s a1 b1 a1ir b1is n o i t) (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_negd a1) (l_e_st_eq_landau_n_rt_rp_posd b1)) => l_e_st_eq_landau_n_rt_rp_r_4r196_t10 r s a1 b1 a1ir b1is n o i t)))))))))). Time Defined. (* constant 4497 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz196f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))), l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_pos s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_pos s))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r196_t11 r s x y t u n o i))))))))). Time Defined. (* constant 4498 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)), l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pnot0 (l_e_st_eq_landau_n_rt_rp_r_ts r s) p) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ts01 r s t)))). Time Defined. (* constant 4499 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t13 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)), l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pnot0 (l_e_st_eq_landau_n_rt_rp_r_ts r s) p) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ts02 r s t)))). Time Defined. (* constant 4500 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t14 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_absp (l_e_st_eq_landau_n_rt_rp_r_ts r s) p) (l_e_st_eq_landau_n_rt_rp_r_satz193 r s)))). Time Defined. (* constant 4501 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz196g : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)), l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_pos s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_neg s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)) => l_e_st_eq_landau_n_rt_rp_r_satz196e r s (l_e_st_eq_landau_n_rt_rp_r_4r196_t12 r s p) (l_e_st_eq_landau_n_rt_rp_r_4r196_t13 r s p) (l_e_st_eq_landau_n_rt_rp_r_4r196_t14 r s p)))). Time Defined. (* constant 4502 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t15 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)), l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_nnot0 (l_e_st_eq_landau_n_rt_rp_r_ts r s) n) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ts01 r s t)))). Time Defined. (* constant 4503 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t16 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)), l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_nnot0 (l_e_st_eq_landau_n_rt_rp_r_ts r s) n) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ts02 r s t)))). Time Defined. (* constant 4504 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r196_t17 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)) => l_e_st_eq_landau_n_rt_rp_r_satz177c (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_satz193a r s) (l_e_st_eq_landau_n_rt_rp_r_absn (l_e_st_eq_landau_n_rt_rp_r_ts r s) n))))). Time Defined. (* constant 4505 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz196h : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)), l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_pos s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)) => l_e_st_eq_landau_n_rt_rp_r_satz196f r s (l_e_st_eq_landau_n_rt_rp_r_4r196_t15 r s n) (l_e_st_eq_landau_n_rt_rp_r_4r196_t16 r s n) (l_e_st_eq_landau_n_rt_rp_r_4r196_t17 r s n)))). Time Defined. (* constant 4506 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r197_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a1) b1) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a1 b1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_satzd197a a1 b1)))))). Time Defined. (* constant 4507 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r197_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_m0d a1) b1) (l_e_st_eq_landau_n_rt_rp_m0d (l_e_st_eq_landau_n_rt_rp_td a1 b1)) (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_m0 r) s (l_e_st_eq_landau_n_rt_rp_m0d a1) b1 (l_e_st_eq_landau_n_rt_rp_r_micm0 r a1 a1ir) b1is) (l_e_st_eq_landau_n_rt_rp_r_micm0 (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is)) (l_e_st_eq_landau_n_rt_rp_r_4r197_t1 r s a1 b1 a1ir b1is))))))). Time Defined. (* constant 4508 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz197a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r197_t2 r s x y t u)))))). Time Defined. (* constant 4509 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz197b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 s) r) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts s r)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_comts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_satz197a s r) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_ts s r) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_comts s r)))). Time Defined. (* constant 4510 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz197c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_satz197a r s) (l_e_st_eq_landau_n_rt_rp_r_satz197b r s))). Time Defined. (* constant 4511 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz197d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_satz197c r s))). Time Defined. (* constant 4512 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz197e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_satz197a r s))). Time Defined. (* constant 4513 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz197f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (l_e_st_eq_landau_n_rt_rp_r_satz197b r s))). Time Defined. (* constant 4514 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz198 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_ts r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s))) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_satz197c r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s)) s r (l_e_st_eq_landau_n_rt_rp_r_satz177 s)))). Time Defined. (* constant 4515 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz198a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_satz198 r s))). Time Defined. (* constant 4516 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr4_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_posd (l_e_st_eq_landau_n_rt_rp_td a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_ptdpp a1 b1 (l_e_st_eq_landau_n_rt_rp_r_posex r a1 a1ir p) (l_e_st_eq_landau_n_rt_rp_r_posex s b1 b1is q))))))))). Time Defined. (* constant 4517 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr4_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_posin (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_ivr4_t3 r s a1 b1 a1ir b1is p q))))))))). Time Defined. (* constant 4518 *) Definition l_e_st_eq_landau_n_rt_rp_r_postspp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_ivr4_t4 r s x y t u p q)))))))). Time Defined. (* constant 4519 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr4_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_negd (l_e_st_eq_landau_n_rt_rp_td a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_ntdpn a1 b1 (l_e_st_eq_landau_n_rt_rp_r_posex r a1 a1ir p) (l_e_st_eq_landau_n_rt_rp_r_negex s b1 b1is n))))))))). Time Defined. (* constant 4520 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr4_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_negin (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_ivr4_t5 r s a1 b1 a1ir b1is p n))))))))). Time Defined. (* constant 4521 *) Definition l_e_st_eq_landau_n_rt_rp_r_negtspn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_ivr4_t6 r s x y t u p n)))))))). Time Defined. (* constant 4522 *) Definition l_e_st_eq_landau_n_rt_rp_r_negtsnp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_isneg (l_e_st_eq_landau_n_rt_rp_r_ts s r) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_comts s r) (l_e_st_eq_landau_n_rt_rp_r_negtspn s r p n))))). Time Defined. (* constant 4523 *) Definition l_e_st_eq_landau_n_rt_rp_r_postsnn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (o:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_e_st_eq_landau_n_rt_rp_r_ispos (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_satz198 r s) (l_e_st_eq_landau_n_rt_rp_r_postspp (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_satz176c r n) (l_e_st_eq_landau_n_rt_rp_r_satz176c s o)))))). Time Defined. (* constant 4524 *) Definition l_e_st_eq_landau_n_rt_rp_r_possq : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_rapp r (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r r)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_postspp r r t t) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts r r)) n) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_postsnn r r t t))). Time Defined. (* constant 4525 *) Definition l_e_st_eq_landau_n_rt_rp_r_nnegsq : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r r))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_0notn (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_satz192a r r t)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_pnotn (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_possq r t))). Time Defined. (* constant 4526 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r199_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_td a2 b2) c2) (l_e_st_eq_landau_n_rt_rp_td a2 (l_e_st_eq_landau_n_rt_rp_td b2 c2))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_satzd199 a2 b2 c2))))))))). Time Defined. (* constant 4527 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r199_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts r s) t) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ts s t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts r s) t) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_td a2 b2) c2) (l_e_st_eq_landau_n_rt_rp_td a2 (l_e_st_eq_landau_n_rt_rp_td b2 c2)) (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_ts r s) t (l_e_st_eq_landau_n_rt_rp_td a2 b2) c2 (l_e_st_eq_landau_n_rt_rp_r_tict r s a2 b2 a2ir b2is) c2it) (l_e_st_eq_landau_n_rt_rp_r_tict r (l_e_st_eq_landau_n_rt_rp_r_ts s t) a2 (l_e_st_eq_landau_n_rt_rp_td b2 c2) a2ir (l_e_st_eq_landau_n_rt_rp_r_tict s t b2 c2 b2is c2it)) (l_e_st_eq_landau_n_rt_rp_r_4r199_t1 r s t a2 b2 c2 a2ir b2is c2it)))))))))). Time Defined. (* constant 4528 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz199 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts r s) t) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ts s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts r s) t) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ts s t))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_4r199_t2 r s t x y z u v w))))))))). Time Defined. (* constant 4529 *) Definition l_e_st_eq_landau_n_rt_rp_r_assts1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts r s) t) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ts s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_satz199 r s t))). Time Defined. (* constant 4530 *) Definition l_e_st_eq_landau_n_rt_rp_r_assts2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts r s) t)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts r s) t) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_satz199 r s t)))). Time Defined. (* constant 4531 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r201_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a2 (l_e_st_eq_landau_n_rt_rp_pd b2 c2)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a2 b2) (l_e_st_eq_landau_n_rt_rp_td a2 c2))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_satzd201 a2 b2 c2))))))))). Time Defined. (* constant 4532 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r201_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t)) (l_e_st_eq_landau_n_rt_rp_td a2 (l_e_st_eq_landau_n_rt_rp_pd b2 c2)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_td a2 b2) (l_e_st_eq_landau_n_rt_rp_td a2 c2)) (l_e_st_eq_landau_n_rt_rp_r_tict r (l_e_st_eq_landau_n_rt_rp_r_pl s t) a2 (l_e_st_eq_landau_n_rt_rp_pd b2 c2) a2ir (l_e_st_eq_landau_n_rt_rp_r_picp s t b2 c2 b2is c2it)) (l_e_st_eq_landau_n_rt_rp_r_picp (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_td a2 b2) (l_e_st_eq_landau_n_rt_rp_td a2 c2) (l_e_st_eq_landau_n_rt_rp_r_tict r s a2 b2 a2ir b2is) (l_e_st_eq_landau_n_rt_rp_r_tict r t a2 c2 a2ir c2it)) (l_e_st_eq_landau_n_rt_rp_r_4r201_t1 r s t a2 b2 c2 a2ir b2is c2it)))))))))). Time Defined. (* constant 4533 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz201 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_4r201_t2 r s t x y z u v w))))))))). Time Defined. (* constant 4534 *) Definition l_e_st_eq_landau_n_rt_rp_r_disttp1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_satz201 t r s) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts t s) (l_e_st_eq_landau_n_rt_rp_r_ts s t) (l_e_st_eq_landau_n_rt_rp_r_comts t r) (l_e_st_eq_landau_n_rt_rp_r_comts t s))))). Time Defined. (* constant 4535 *) Definition l_e_st_eq_landau_n_rt_rp_r_disttp2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_satz201 r s t))). Time Defined. (* constant 4536 *) Definition l_e_st_eq_landau_n_rt_rp_r_distpt1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_pl r s) t)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_pl r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_disttp1 r s t)))). Time Defined. (* constant 4537 *) Definition l_e_st_eq_landau_n_rt_rp_r_distpt2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t)) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_pl s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_pl s t)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t)) (l_e_st_eq_landau_n_rt_rp_r_disttp2 r s t)))). Time Defined. (* constant 4538 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz202 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_mn s t)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_mn s t)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 t))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t)) (l_e_st_eq_landau_n_rt_rp_r_disttp2 r s (l_e_st_eq_landau_n_rt_rp_r_m0 t)) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_m0 t)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts r t)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_satz197b r t))))). Time Defined. (* constant 4539 *) Definition l_e_st_eq_landau_n_rt_rp_r_disttm1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_mn r s) t) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_mn r s) t) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 s) t)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_disttp1 r (l_e_st_eq_landau_n_rt_rp_r_m0 s) t) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 s) t) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_satz197a s t))))). Time Defined. (* constant 4540 *) Definition l_e_st_eq_landau_n_rt_rp_r_disttm2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_mn s t)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_satz202 r s t))). Time Defined. (* constant 4541 *) Definition l_e_st_eq_landau_n_rt_rp_r_distmt1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_mn r s) t)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_mn r s) t) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (l_e_st_eq_landau_n_rt_rp_r_disttm1 r s t)))). Time Defined. (* constant 4542 *) Definition l_e_st_eq_landau_n_rt_rp_r_distmt2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t)) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_mn s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_mn s t)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t)) (l_e_st_eq_landau_n_rt_rp_r_disttm2 r s t)))). Time Defined. (* constant 4543 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz200 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_mn s t)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_satz202 r s t))). Time Defined. (* constant 4544 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r203_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos t), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_td a2 c2) (l_e_st_eq_landau_n_rt_rp_td b2 c2)))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos t) => l_e_st_eq_landau_n_rt_rp_satzd203a a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_moreex r s a2 b2 a2ir b2is m) (l_e_st_eq_landau_n_rt_rp_r_posex t c2 c2it p)))))))))))). Time Defined. (* constant 4545 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r203_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos t), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos t) => l_e_st_eq_landau_n_rt_rp_r_morein (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t) (l_e_st_eq_landau_n_rt_rp_td a2 c2) (l_e_st_eq_landau_n_rt_rp_td b2 c2) (l_e_st_eq_landau_n_rt_rp_r_tict r t a2 c2 a2ir c2it) (l_e_st_eq_landau_n_rt_rp_r_tict s t b2 c2 b2is c2it) (l_e_st_eq_landau_n_rt_rp_r_4r203_t1 r s t a2 b2 c2 a2ir b2is c2it m p)))))))))))). Time Defined. (* constant 4546 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos t), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos t) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_4r203_t2 r s t x y z u v w m p))))))))))). Time Defined. (* constant 4547 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is t l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is t l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts02 r t i) (l_e_st_eq_landau_n_rt_rp_r_ts02 s t i)))))). Time Defined. (* constant 4548 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r203_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg t), l_e_st_eq_landau_n_rt_rp_lessd (l_e_st_eq_landau_n_rt_rp_td a2 c2) (l_e_st_eq_landau_n_rt_rp_td b2 c2)))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg t) => l_e_st_eq_landau_n_rt_rp_satzd203c a2 b2 c2 (l_e_st_eq_landau_n_rt_rp_r_moreex r s a2 b2 a2ir b2is m) (l_e_st_eq_landau_n_rt_rp_r_negex t c2 c2it n)))))))))))). Time Defined. (* constant 4549 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r203_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a2:l_e_st_eq_landau_n_rt_rp_dif), (forall (b2:l_e_st_eq_landau_n_rt_rp_dif), (forall (c2:l_e_st_eq_landau_n_rt_rp_dif), (forall (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg t), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c2:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a2ir:l_e_st_eq_landau_n_rt_rp_r_inn a2 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b2is:l_e_st_eq_landau_n_rt_rp_r_inn b2 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c2it:l_e_st_eq_landau_n_rt_rp_r_inn c2 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg t) => l_e_st_eq_landau_n_rt_rp_r_lessin (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t) (l_e_st_eq_landau_n_rt_rp_td a2 c2) (l_e_st_eq_landau_n_rt_rp_td b2 c2) (l_e_st_eq_landau_n_rt_rp_r_tict r t a2 c2 a2ir c2it) (l_e_st_eq_landau_n_rt_rp_r_tict s t b2 c2 b2is c2it) (l_e_st_eq_landau_n_rt_rp_r_4r203_t3 r s t a2 b2 c2 a2ir b2is c2it m n)))))))))))). Time Defined. (* constant 4550 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg t), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg t) => l_e_st_eq_landau_n_rt_rp_r_realapp3 r s t (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_4r203_t4 r s t x y z u v w m n))))))))))). Time Defined. (* constant 4551 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos t), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos t) => l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts s t) (l_e_st_eq_landau_n_rt_rp_r_ts t s) (l_e_st_eq_landau_n_rt_rp_r_comts r t) (l_e_st_eq_landau_n_rt_rp_r_comts s t) (l_e_st_eq_landau_n_rt_rp_r_satz203a r s t m p)))))). Time Defined. (* constant 4552 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is t l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is t l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 t r i) (l_e_st_eq_landau_n_rt_rp_r_ts01 t s i)))))). Time Defined. (* constant 4553 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg t), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg t) => l_e_st_eq_landau_n_rt_rp_r_isless12 (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts s t) (l_e_st_eq_landau_n_rt_rp_r_ts t s) (l_e_st_eq_landau_n_rt_rp_r_comts r t) (l_e_st_eq_landau_n_rt_rp_r_comts s t) (l_e_st_eq_landau_n_rt_rp_r_satz203c r s t m n)))))). Time Defined. (* constant 4554 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203g : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos t), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos t) => l_e_st_eq_landau_n_rt_rp_r_lemma1 (l_e_st_eq_landau_n_rt_rp_r_ts s t) (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_satz203a s r t (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l) p)))))). Time Defined. (* constant 4555 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203h : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is t l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is t l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts02 r t i) (l_e_st_eq_landau_n_rt_rp_r_ts02 s t i)))))). Time Defined. (* constant 4556 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203j : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg t), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_ts s t)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg t) => l_e_st_eq_landau_n_rt_rp_r_lemma2 (l_e_st_eq_landau_n_rt_rp_r_ts s t) (l_e_st_eq_landau_n_rt_rp_r_ts r t) (l_e_st_eq_landau_n_rt_rp_r_satz203c s r t (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l) n)))))). Time Defined. (* constant 4557 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203k : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos t), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos t) => l_e_st_eq_landau_n_rt_rp_r_lemma1 (l_e_st_eq_landau_n_rt_rp_r_ts t s) (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_satz203d s r t (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l) p)))))). Time Defined. (* constant 4558 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203l : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is t l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is t l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 t r i) (l_e_st_eq_landau_n_rt_rp_r_ts01 t s i)))))). Time Defined. (* constant 4559 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz203m : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg t), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_ts t s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg t) => l_e_st_eq_landau_n_rt_rp_r_lemma2 (l_e_st_eq_landau_n_rt_rp_r_ts t s) (l_e_st_eq_landau_n_rt_rp_r_ts t r) (l_e_st_eq_landau_n_rt_rp_r_satz203f s r t (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l) n)))))). Time Defined. (* constant 4560 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0), l_not (l_e_st_eq_landau_n_rt_rp_zero a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_zero a0) (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) n1 (fun (t:l_e_st_eq_landau_n_rt_rp_zero a0) => l_e_st_eq_landau_n_rt_rp_r_0in r a0 a0ir t))))). Time Defined. (* constant 4561 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a3:l_e_st_eq_landau_n_rt_rp_dif), (forall (b3:l_e_st_eq_landau_n_rt_rp_dif), (forall (c3:l_e_st_eq_landau_n_rt_rp_dif), (forall (d3:l_e_st_eq_landau_n_rt_rp_dif), (forall (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s t) r), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s u) r), l_e_st_eq_landau_n_rt_rp_eq c3 d3))))))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s t) r) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s u) r) => l_e_st_eq_landau_n_rt_rp_satzd204b a3 b3 (l_e_st_eq_landau_n_rt_rp_r_4r204_t1 s b3 b3is n1) c3 d3 (l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_ts s t) r (l_e_st_eq_landau_n_rt_rp_td b3 c3) a3 (l_e_st_eq_landau_n_rt_rp_r_tict s t b3 c3 b3is c3it) a3ir i) (l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_ts s u) r (l_e_st_eq_landau_n_rt_rp_td b3 d3) a3 (l_e_st_eq_landau_n_rt_rp_r_tict s u b3 d3 b3is d3iu) a3ir j)))))))))))))))). Time Defined. (* constant 4562 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a3:l_e_st_eq_landau_n_rt_rp_dif), (forall (b3:l_e_st_eq_landau_n_rt_rp_dif), (forall (c3:l_e_st_eq_landau_n_rt_rp_dif), (forall (d3:l_e_st_eq_landau_n_rt_rp_dif), (forall (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s t) r), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s u) r), l_e_st_eq_landau_n_rt_rp_r_is t u))))))))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (c3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (d3:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a3ir:l_e_st_eq_landau_n_rt_rp_r_inn a3 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b3is:l_e_st_eq_landau_n_rt_rp_r_inn b3 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (c3it:l_e_st_eq_landau_n_rt_rp_r_inn c3 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (d3iu:l_e_st_eq_landau_n_rt_rp_r_inn d3 (l_e_st_eq_landau_n_rt_rp_r_class u)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s t) r) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s u) r) => l_e_st_eq_landau_n_rt_rp_r_isin t u c3 d3 c3it d3iu (l_e_st_eq_landau_n_rt_rp_r_4r204_t2 r s t u a3 b3 c3 d3 a3ir b3is c3it d3iu n1 i j)))))))))))))))). Time Defined. (* constant 4563 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz204b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s y) r), l_e_st_eq_landau_n_rt_rp_r_is x y))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s y) r) => l_e_st_eq_landau_n_rt_rp_r_realapp4 r s x y (l_e_st_eq_landau_n_rt_rp_r_is x y) (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_dif) => (fun (v:l_e_st_eq_landau_n_rt_rp_dif) => (fun (w:l_e_st_eq_landau_n_rt_rp_dif) => (fun (zi:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (ui:l_e_st_eq_landau_n_rt_rp_r_inn u (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (vi:l_e_st_eq_landau_n_rt_rp_r_inn v (l_e_st_eq_landau_n_rt_rp_r_class x)) => (fun (wi:l_e_st_eq_landau_n_rt_rp_r_inn w (l_e_st_eq_landau_n_rt_rp_r_class y)) => l_e_st_eq_landau_n_rt_rp_r_4r204_t3 r s x y z u v w zi ui vi wi n i j))))))))))))))). Time Defined. (* constant 4564 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_some l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 x) a1)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_satzd204a a1 b1 (l_e_st_eq_landau_n_rt_rp_r_4r204_t1 s b1 b1is n1)))))))). Time Defined. (* constant 4565 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_ar : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 a) a1), l_e_st_eq_landau_n_rt_rp_r_real))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 a) a1) => l_e_st_eq_landau_n_rt_rp_r_realof a))))))))). Time Defined. (* constant 4566 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 a) a1), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_4r204_ar r s a1 b1 a1ir b1is n1 a e)) r))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 a) a1) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_4r204_ar r s a1 b1 a1ir b1is n1 a e)) r (l_e_st_eq_landau_n_rt_rp_td b1 a) a1 (l_e_st_eq_landau_n_rt_rp_r_tict s (l_e_st_eq_landau_n_rt_rp_r_4r204_ar r s a1 b1 a1ir b1is n1 a e) b1 a b1is (l_e_st_eq_landau_n_rt_rp_r_innclass a)) a1ir e))))))))). Time Defined. (* constant 4567 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (a:l_e_st_eq_landau_n_rt_rp_dif), (forall (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 a) a1), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r)))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => (fun (e:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 a) a1) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r) (l_e_st_eq_landau_n_rt_rp_r_4r204_ar r s a1 b1 a1ir b1is n1 a e) (l_e_st_eq_landau_n_rt_rp_r_4r204_t5 r s a1 b1 a1ir b1is n1 a e)))))))))). Time Defined. (* constant 4568 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 x) a1) (l_e_st_eq_landau_n_rt_rp_r_4r204_t4 r s a1 b1 a1ir b1is n1) (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b1 x) a1) => l_e_st_eq_landau_n_rt_rp_r_4r204_t6 r s a1 b1 a1ir b1is n1 x t))))))))). Time Defined. (* constant 4569 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz204a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_4r204_t7 r s x y t u n))))))). Time Defined. (* constant 4570 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz204 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_one (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_onei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s y) r) => l_e_st_eq_landau_n_rt_rp_r_satz204b r s n x y t u)))) (l_e_st_eq_landau_n_rt_rp_r_satz204a r s n)))). Time Defined. (* constant 4571 *) Definition l_e_st_eq_landau_n_rt_rp_r_ov : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_ind l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r) (l_e_st_eq_landau_n_rt_rp_r_satz204 r s n)))). Time Defined. (* constant 4572 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz204c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_oneax l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r) (l_e_st_eq_landau_n_rt_rp_r_satz204 r s n)))). Time Defined. (* constant 4573 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz204d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_ov r s n))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) r (l_e_st_eq_landau_n_rt_rp_r_satz204c r s n)))). Time Defined. (* constant 4574 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz204e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ov r s n) s) r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ov r s n) s) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) r (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_ov r s n) s) (l_e_st_eq_landau_n_rt_rp_r_satz204c r s n)))). Time Defined. (* constant 4575 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz204f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ov r s n) s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ov r s n) s) r (l_e_st_eq_landau_n_rt_rp_r_satz204e r s n)))). Time Defined. (* constant 4576 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz204g : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r), l_e_st_eq_landau_n_rt_rp_r_is x (l_e_st_eq_landau_n_rt_rp_r_ov r s n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s x) r) => l_e_st_eq_landau_n_rt_rp_r_satz204b r s n x (l_e_st_eq_landau_n_rt_rp_r_ov r s n) i (l_e_st_eq_landau_n_rt_rp_r_satz204c r s n)))))). Time Defined. (* constant 4577 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_ros : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ov r s n))). Time Defined. (* constant 4578 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_ispos r (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)) (l_e_st_eq_landau_n_rt_rp_r_satz204d r s n) p)))). Time Defined. (* constant 4579 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_and_th1 (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)) (l_e_st_eq_landau_n_rt_rp_r_pnotn s q)))))). Time Defined. (* constant 4580 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_and (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_ore1 (l_and (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))) (l_e_st_eq_landau_n_rt_rp_r_satz196g s (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n) (l_e_st_eq_landau_n_rt_rp_r_4r204_t8 r s n p)) (l_e_st_eq_landau_n_rt_rp_r_4r204_t9 r s n p q)))))). Time Defined. (* constant 4581 *) Definition l_e_st_eq_landau_n_rt_rp_r_posovpp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ov r s n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) (l_e_st_eq_landau_n_rt_rp_r_4r204_t10 r s n p q)))))). Time Defined. (* constant 4582 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg s), l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_and_th1 (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)) (l_e_st_eq_landau_n_rt_rp_r_nnotp s m)))))). Time Defined. (* constant 4583 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg s), l_and (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_ore2 (l_and (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))) (l_e_st_eq_landau_n_rt_rp_r_satz196g s (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n) (l_e_st_eq_landau_n_rt_rp_r_4r204_t8 r s n p)) (l_e_st_eq_landau_n_rt_rp_r_4r204_t11 r s n p m)))))). Time Defined. (* constant 4584 *) Definition l_e_st_eq_landau_n_rt_rp_r_negovpn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ov r s n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) (l_e_st_eq_landau_n_rt_rp_r_4r204_t12 r s n p m)))))). Time Defined. (* constant 4585 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t13 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_isneg r (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)) (l_e_st_eq_landau_n_rt_rp_r_satz204d r s n) m)))). Time Defined. (* constant 4586 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t14 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_and_th1 (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)) (l_e_st_eq_landau_n_rt_rp_r_pnotn s p)))))). Time Defined. (* constant 4587 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t15 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_and (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_ore1 (l_and (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))) (l_e_st_eq_landau_n_rt_rp_r_satz196h s (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n) (l_e_st_eq_landau_n_rt_rp_r_4r204_t13 r s n m)) (l_e_st_eq_landau_n_rt_rp_r_4r204_t14 r s n m p)))))). Time Defined. (* constant 4588 *) Definition l_e_st_eq_landau_n_rt_rp_r_negovnp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ov r s n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) (l_e_st_eq_landau_n_rt_rp_r_4r204_t15 r s n m p)))))). Time Defined. (* constant 4589 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t16 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (l:l_e_st_eq_landau_n_rt_rp_r_neg s), l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_and_th1 (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n)) (l_e_st_eq_landau_n_rt_rp_r_nnotp s l)))))). Time Defined. (* constant 4590 *) Definition l_e_st_eq_landau_n_rt_rp_r_4r204_t17 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (l:l_e_st_eq_landau_n_rt_rp_r_neg s), l_and (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_ore2 (l_and (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n))) (l_e_st_eq_landau_n_rt_rp_r_satz196h s (l_e_st_eq_landau_n_rt_rp_r_4r204_ros r s n) (l_e_st_eq_landau_n_rt_rp_r_4r204_t13 r s n m)) (l_e_st_eq_landau_n_rt_rp_r_4r204_t16 r s n m l)))))). Time Defined. (* constant 4591 *) Definition l_e_st_eq_landau_n_rt_rp_r_posovnn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (m:l_e_st_eq_landau_n_rt_rp_r_neg r), (forall (l:l_e_st_eq_landau_n_rt_rp_r_neg s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ov r s n)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_neg r) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_neg s) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) (l_e_st_eq_landau_n_rt_rp_r_4r204_t17 r s n m l)))))). Time Defined. (* constant 4592 *) Definition l_e_st_eq_landau_n_rt_rp_r_morerpep : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r0 s0), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r0 s0) => l_e_st_eq_landau_n_rt_rp_r_morein (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp s0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp s0)) (l_e_st_eq_landau_n_rt_rp_morerpepd r0 s0 m)))). Time Defined. (* constant 4593 *) Definition l_e_st_eq_landau_n_rt_rp_r_morerpip : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)), l_e_st_eq_landau_n_rt_rp_more r0 s0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_morerpipd r0 s0 (l_e_st_eq_landau_n_rt_rp_r_moreex (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp s0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp s0)) m)))). Time Defined. (* constant 4594 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessrpep : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_less r0 s0), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_less r0 s0) => l_e_st_eq_landau_n_rt_rp_r_lemma1 (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_morerpep s0 r0 (l_e_st_eq_landau_n_rt_rp_satz122 r0 s0 l))))). Time Defined. (* constant 4595 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessrpip : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)), l_e_st_eq_landau_n_rt_rp_less r0 s0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_satz121 s0 r0 (l_e_st_eq_landau_n_rt_rp_r_morerpip s0 r0 (l_e_st_eq_landau_n_rt_rp_r_lemma2 (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) l))))). Time Defined. (* constant 4596 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr5_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp r p)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp s q))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_ismore12 r (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp r p)) s (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)) (l_e_st_eq_landau_n_rt_rp_r_isprp1 r p) (l_e_st_eq_landau_n_rt_rp_r_isprp1 s q) m))))). Time Defined. (* constant 4597 *) Definition l_e_st_eq_landau_n_rt_rp_r_moreperp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r s), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r s) => l_e_st_eq_landau_n_rt_rp_r_morerpip (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q) (l_e_st_eq_landau_n_rt_rp_r_ivr5_t1 r s p q m)))))). Time Defined. (* constant 4598 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr5_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp r p)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp s q))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)) => l_e_st_eq_landau_n_rt_rp_r_morerpep (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q) m))))). Time Defined. (* constant 4599 *) Definition l_e_st_eq_landau_n_rt_rp_r_morepirp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)), l_e_st_eq_landau_n_rt_rp_r_more r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (m:l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)) => l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp r p)) r (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)) s (l_e_st_eq_landau_n_rt_rp_r_isprp2 r p) (l_e_st_eq_landau_n_rt_rp_r_isprp2 s q) (l_e_st_eq_landau_n_rt_rp_r_ivr5_t2 r s p q m)))))). Time Defined. (* constant 4600 *) Definition l_e_st_eq_landau_n_rt_rp_r_lessperp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_satz121 (l_e_st_eq_landau_n_rt_rp_r_rpofp s q) (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_moreperp s r q p (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l))))))). Time Defined. (* constant 4601 *) Definition l_e_st_eq_landau_n_rt_rp_r_lesspirp : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos s), (forall (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)), l_e_st_eq_landau_n_rt_rp_r_less r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos s) => (fun (l:l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q)) => l_e_st_eq_landau_n_rt_rp_r_lemma1 s r (l_e_st_eq_landau_n_rt_rp_r_morepirp s r q p (l_e_st_eq_landau_n_rt_rp_satz122 (l_e_st_eq_landau_n_rt_rp_r_rpofp r p) (l_e_st_eq_landau_n_rt_rp_r_rpofp s q) l))))))). Time Defined. (* constant 4602 *) Definition l_e_st_eq_landau_n_rt_rp_r_s01 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_setof l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_lessis x r)). Time Defined. (* constant 4603 *) Definition l_e_st_eq_landau_n_rt_rp_r_s02 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_setof l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_more x r)). Time Defined. (* constant 4604 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s01 r))), l_not (l_e_st_eq_landau_n_rt_rp_r_lessis s r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s01 r))) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_lessis s r) (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s01 r)) n (fun (t:l_e_st_eq_landau_n_rt_rp_r_lessis s r) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_lessis x r) s t)))). Time Defined. (* constant 4605 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s01 r))), l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s02 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s01 r))) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_more x r) s (l_e_st_eq_landau_n_rt_rp_r_satz167k s r (l_e_st_eq_landau_n_rt_rp_r_5r205_t1 r s n))))). Time Defined. (* constant 4606 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb00 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s01 r)) (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s02 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s01 r))) => l_e_st_eq_landau_n_rt_rp_r_5r205_t2 r x t))). Time Defined. (* constant 4607 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_s01 r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_lessis x r) r (l_e_st_eq_landau_n_rt_rp_r_lessisi2 r r (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r))). Time Defined. (* constant 4608 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb01a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_s01 r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_nonemptyi l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_s01 r) r (l_e_st_eq_landau_n_rt_rp_r_5r205_t3 r)). Time Defined. (* constant 4609 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_1rl) r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ismore2 (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_0) r (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl02 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_satz188k l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0 r (l_e_st_eq_landau_n_rt_rp_r_satz169a l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_pos1))). Time Defined. (* constant 4610 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_s02 r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_more x r) (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_5r205_t4 r)). Time Defined. (* constant 4611 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb01b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_s02 r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_nonemptyi l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_s02 r) (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_5r205_t5 r)). Time Defined. (* constant 4612 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s01 r)), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_in t (l_e_st_eq_landau_n_rt_rp_r_s02 r)), l_e_st_eq_landau_n_rt_rp_r_less s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s01 r)) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_in t (l_e_st_eq_landau_n_rt_rp_r_s02 r)) => l_e_st_eq_landau_n_rt_rp_r_satz172a s r t (l_e_st_estie l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_lessis x r) s i) (l_e_st_eq_landau_n_rt_rp_r_lemma1 t r (l_e_st_estie l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_more x r) t j))))))). Time Defined. (* constant 4613 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb02 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s01 r)), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y (l_e_st_eq_landau_n_rt_rp_r_s02 r)), l_e_st_eq_landau_n_rt_rp_r_less x y))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s01 r)) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_in y (l_e_st_eq_landau_n_rt_rp_r_s02 r)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t6 r x t y u))))). Time Defined. (* constant 4614 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb03a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s r), l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s01 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s r) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_lessis x r) s (l_e_st_eq_landau_n_rt_rp_r_lessisi1 s r l)))). Time Defined. (* constant 4615 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb03b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s r), l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s02 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s r) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_more x r) s m))). Time Defined. (* constant 4616 *) Definition l_e_st_eq_landau_n_rt_rp_r_s11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_setof l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_less x r)). Time Defined. (* constant 4617 *) Definition l_e_st_eq_landau_n_rt_rp_r_s12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_setof l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_moreis x r)). Time Defined. (* constant 4618 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s11 r))), l_not (l_e_st_eq_landau_n_rt_rp_r_less s r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s11 r))) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_less s r) (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s11 r)) n (fun (t:l_e_st_eq_landau_n_rt_rp_r_less s r) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_less x r) s t)))). Time Defined. (* constant 4619 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s11 r))), l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s12 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s11 r))) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_moreis x r) s (l_e_st_eq_landau_n_rt_rp_r_satz167f s r (l_e_st_eq_landau_n_rt_rp_r_5r205_t7 r s n))))). Time Defined. (* constant 4620 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s11 r)) (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s12 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s11 r))) => l_e_st_eq_landau_n_rt_rp_r_5r205_t8 r x t))). Time Defined. (* constant 4621 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_mn r l_e_st_eq_landau_n_rt_rp_r_1rl) r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_isless2 (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_0) r (l_e_st_eq_landau_n_rt_rp_r_mn r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl02 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_satz188m (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) l_e_st_eq_landau_n_rt_rp_r_0 r (l_e_st_eq_landau_n_rt_rp_r_satz169c (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_satz176a l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_pos1)))). Time Defined. (* constant 4622 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_mn r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_s11 r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_less x r) (l_e_st_eq_landau_n_rt_rp_r_mn r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_5r205_t9 r)). Time Defined. (* constant 4623 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb11a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_s11 r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_nonemptyi l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_s11 r) (l_e_st_eq_landau_n_rt_rp_r_mn r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_5r205_t10 r)). Time Defined. (* constant 4624 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_s12 r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_moreis x r) r (l_e_st_eq_landau_n_rt_rp_r_moreisi2 r r (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real r))). Time Defined. (* constant 4625 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb11b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_s12 r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_nonemptyi l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_s12 r) r (l_e_st_eq_landau_n_rt_rp_r_5r205_t11 r)). Time Defined. (* constant 4626 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s11 r)), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_in t (l_e_st_eq_landau_n_rt_rp_r_s12 r)), l_e_st_eq_landau_n_rt_rp_r_less s t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s11 r)) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_in t (l_e_st_eq_landau_n_rt_rp_r_s12 r)) => l_e_st_eq_landau_n_rt_rp_r_satz172b s r t (l_e_st_estie l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_less x r) s i) (l_e_st_eq_landau_n_rt_rp_r_satz168a t r (l_e_st_estie l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_moreis x r) t j))))))). Time Defined. (* constant 4627 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s11 r)), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y (l_e_st_eq_landau_n_rt_rp_r_s12 r)), l_e_st_eq_landau_n_rt_rp_r_less x y))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_s11 r)) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_in y (l_e_st_eq_landau_n_rt_rp_r_s12 r)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t12 r x t y u))))). Time Defined. (* constant 4628 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb13a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s r), l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s11 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s r) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_less x r) s l))). Time Defined. (* constant 4629 *) Definition l_e_st_eq_landau_n_rt_rp_r_vb13b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s r), l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_s12 r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s r) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_moreis x r) s (l_e_st_eq_landau_n_rt_rp_r_moreisi1 s r m)))). Time Defined. (* constant 4630 *) Definition l_e_st_eq_landau_n_rt_rp_r_2rl : l_e_st_eq_landau_n_rt_rp_r_real. exact (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl). Time Defined. (* constant 4631 *) Definition l_e_st_eq_landau_n_rt_rp_r_pos2 : l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_2rl. exact (l_e_st_eq_landau_n_rt_rp_r_pospl l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_pos1 l_e_st_eq_landau_n_rt_rp_r_pos1). Time Defined. (* constant 4632 *) Definition l_e_st_eq_landau_n_rt_rp_r_half : l_e_st_eq_landau_n_rt_rp_r_real. exact (l_e_st_eq_landau_n_rt_rp_r_ov l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_2rl (l_e_st_eq_landau_n_rt_rp_r_pnot0 l_e_st_eq_landau_n_rt_rp_r_2rl l_e_st_eq_landau_n_rt_rp_r_pos2)). Time Defined. (* constant 4633 *) Definition l_e_st_eq_landau_n_rt_rp_r_poshalf : l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_half. exact (l_e_st_eq_landau_n_rt_rp_r_posovpp l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_2rl (l_e_st_eq_landau_n_rt_rp_r_pnot0 l_e_st_eq_landau_n_rt_rp_r_2rl l_e_st_eq_landau_n_rt_rp_r_pos2) l_e_st_eq_landau_n_rt_rp_r_pos1 l_e_st_eq_landau_n_rt_rp_r_pos2). Time Defined. (* constant 4634 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr5_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl r r) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_2rl r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl r r) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_2rl r) (l_e_st_eq_landau_n_rt_rp_r_ispl12 r (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r) r (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_satz195c r) (l_e_st_eq_landau_n_rt_rp_r_satz195c r)) (l_e_st_eq_landau_n_rt_rp_r_distpt1 l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl r)). Time Defined. (* constant 4635 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr5_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r r)) r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r r)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_2rl r)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half l_e_st_eq_landau_n_rt_rp_r_2rl) r) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl r) r (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_pl r r) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_2rl r) l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_ivr5_t3 r)) (l_e_st_eq_landau_n_rt_rp_r_assts2 l_e_st_eq_landau_n_rt_rp_r_half l_e_st_eq_landau_n_rt_rp_r_2rl r) (l_e_st_eq_landau_n_rt_rp_r_ists1 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half l_e_st_eq_landau_n_rt_rp_r_2rl) l_e_st_eq_landau_n_rt_rp_r_1rl r (l_e_st_eq_landau_n_rt_rp_r_satz204e l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_2rl (l_e_st_eq_landau_n_rt_rp_r_pnot0 l_e_st_eq_landau_n_rt_rp_r_2rl l_e_st_eq_landau_n_rt_rp_r_pos2))) (l_e_st_eq_landau_n_rt_rp_r_satz195b r)). Time Defined. (* constant 4636 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr5_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r r)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_satz203k (l_e_st_eq_landau_n_rt_rp_r_pl r r) (l_e_st_eq_landau_n_rt_rp_r_pl r s) l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_satz188m r s r l) l_e_st_eq_landau_n_rt_rp_r_poshalf))). Time Defined. (* constant 4637 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_less r (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_isless1 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r r)) r (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_ivr5_t4 r) (l_e_st_eq_landau_n_rt_rp_r_ivr5_t5 r s l)))). Time Defined. (* constant 4638 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr5_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl s s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_satz203k (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl s s) l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_satz188f r s s l) l_e_st_eq_landau_n_rt_rp_r_poshalf))). Time Defined. (* constant 4639 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r s)) s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_isless2 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl s s)) s (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (l_e_st_eq_landau_n_rt_rp_r_ivr5_t4 s) (l_e_st_eq_landau_n_rt_rp_r_ivr5_t6 r s l)))). Time Defined. (* constant 4640 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_pos s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_satz169b s (l_e_st_eq_landau_n_rt_rp_r_trmore s r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l) (l_e_st_eq_landau_n_rt_rp_r_satz169a r p)))))). Time Defined. (* constant 4641 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_less x r), l_e_st_eq_landau_n_rt_rp_r_in x s1))))). Time Defined. (* constant 4642 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_more x r), l_e_st_eq_landau_n_rt_rp_r_in x s2))))). Time Defined. (* constant 4643 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 s1 s2 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 s1 s2 r)))). Time Defined. (* constant 4644 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_mxy : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less x y), l_e_st_eq_landau_n_rt_rp_r_real))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less x y) => l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_half (l_e_st_eq_landau_n_rt_rp_r_pl x y)))))))))))). Time Defined. (* constant 4645 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t13 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less x y), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) x))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less x y) => l_e_st_eq_landau_n_rt_rp_r_lemma2 x (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_lemma3 x y l)))))))))))). Time Defined. (* constant 4646 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t14 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less x y), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) y))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less x y) => l_e_st_eq_landau_n_rt_rp_r_lemma4 x y l))))))))))). Time Defined. (* constant 4647 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t15 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less x y), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) s1))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less x y) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 s1 s2 y) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 s1 s2 y) py (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_t14 s1 s2 p0 p1a p1b p2 x y px py l)))))))))))). Time Defined. (* constant 4648 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t16 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less x y), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) s2))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less x y) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 s1 s2 x) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 s1 s2 x) px (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_t13 s1 s2 p0 p1a p1b p2 x y px py l)))))))))))). Time Defined. (* constant 4649 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t17 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less x y), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less x y) => p2 (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_t15 s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_t16 s1 s2 p0 p1a p1b p2 x y px py l)))))))))))). Time Defined. (* constant 4650 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t18 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less x y), l_con))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less x y) => l_ec3e31 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l)) (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l)) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l)) (l_e_st_eq_landau_n_rt_rp_r_satz167b (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l) (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l)) (l_e_st_eq_landau_n_rt_rp_r_5r205_t17 s1 s2 p0 p1a p1b p2 x y px py l) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_mxy s1 s2 p0 p1a p1b p2 x y px py l))))))))))))). Time Defined. (* constant 4651 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t19 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), l_not (l_e_st_eq_landau_n_rt_rp_r_less x y))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_less x y) => l_e_st_eq_landau_n_rt_rp_r_5r205_t18 s1 s2 p0 p1a p1b p2 x y px py t))))))))))). Time Defined. (* constant 4652 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t20 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), l_not (l_e_st_eq_landau_n_rt_rp_r_more x y))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_more x y) => l_e_st_eq_landau_n_rt_rp_r_5r205_t18 s1 s2 p0 p1a p1b p2 y x py px (l_e_st_eq_landau_n_rt_rp_r_lemma1 x y t)))))))))))). Time Defined. (* constant 4653 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t21 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x), (forall (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y), l_e_st_eq_landau_n_rt_rp_r_is x y)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (px:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (py:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => l_or3e1 (l_e_st_eq_landau_n_rt_rp_r_is x y) (l_e_st_eq_landau_n_rt_rp_r_more x y) (l_e_st_eq_landau_n_rt_rp_r_less x y) (l_e_st_eq_landau_n_rt_rp_r_satz167a x y) (l_e_st_eq_landau_n_rt_rp_r_5r205_t20 s1 s2 p0 p1a p1b p2 x y px py) (l_e_st_eq_landau_n_rt_rp_r_5r205_t19 s1 s2 p0 p1a p1b p2 x y px py))))))))))). Time Defined. (* constant 4654 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t22 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), l_e_amone l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 y) => l_e_st_eq_landau_n_rt_rp_r_5r205_t21 s1 s2 p0 p1a p1b p2 x y t u)))))))))). Time Defined. (* constant 4655 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t23 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_r_pos r))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1) a))))))))). Time Defined. (* constant 4656 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t24 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_r_in r s1))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1) a))))))))). Time Defined. (* constant 4657 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_set l_e_st_eq_landau_n_rt_cut))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_setof l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp x) s1)))))))))). Time Defined. (* constant 4658 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_set l_e_st_eq_landau_n_rt_cut))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_setof l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp x) s2)))))))))). Time Defined. (* constant 4659 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t25 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s1), l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s1) => l_e_st_estii l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp x) s1) r0 i))))))))))). Time Defined. (* constant 4660 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t26 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s1))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_estie l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp x) s1) r0 i))))))))))). Time Defined. (* constant 4661 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t27 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s2), l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s2) => l_e_st_estii l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp x) s2) r0 i))))))))))). Time Defined. (* constant 4662 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t28 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s2))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_estie l_e_st_eq_landau_n_rt_cut (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp x) s2) r0 i))))))))))). Time Defined. (* constant 4663 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t29 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (r0:l_e_st_eq_landau_n_rt_cut), l_or (l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)) (l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s1) (l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s2) (l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)) (l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)) (p0 (l_e_st_eq_landau_n_rt_rp_r_pofrp r0)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s1) => l_e_st_eq_landau_n_rt_rp_r_5r205_t25 s1 s2 p0 p1a p1b p2 case1 r a r0 t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) s2) => l_e_st_eq_landau_n_rt_rp_r_5r205_t27 s1 s2 p0 p1a p1b p2 case1 r a r0 t))))))))))). Time Defined. (* constant 4664 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_pr1 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_cut))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_eq_landau_n_rt_rp_r_rpofp r (l_e_st_eq_landau_n_rt_rp_r_5r205_t23 s1 s2 p0 p1a p1b p2 case1 r a)))))))))). Time Defined. (* constant 4665 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t30 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_pr1 s1 s2 p0 p1a p1b p2 case1 r a)) s1))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s1) r (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_pr1 s1 s2 p0 p1a p1b p2 case1 r a)) (l_e_st_eq_landau_n_rt_rp_r_5r205_t24 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_isprp1 r (l_e_st_eq_landau_n_rt_rp_r_5r205_t23 s1 s2 p0 p1a p1b p2 case1 r a))))))))))). Time Defined. (* constant 4666 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t31 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_nonempty l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_nonemptyi l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_pr1 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_t25 s1 s2 p0 p1a p1b p2 case1 r a (l_e_st_eq_landau_n_rt_rp_r_5r205_pr1 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_t30 s1 s2 p0 p1a p1b p2 case1 r a))))))))))). Time Defined. (* constant 4667 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t32 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s s2), l_e_st_eq_landau_n_rt_rp_r_less r s))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s s2) => p2 r (l_e_st_eq_landau_n_rt_rp_r_5r205_t24 s1 s2 p0 p1a p1b p2 case1 r a) s i))))))))))). Time Defined. (* constant 4668 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t33 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s s2), l_e_st_eq_landau_n_rt_rp_r_pos s))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s s2) => l_e_st_eq_landau_n_rt_rp_r_lemma5 r s (l_e_st_eq_landau_n_rt_rp_r_5r205_t32 s1 s2 p0 p1a p1b p2 case1 r a s i) (l_e_st_eq_landau_n_rt_rp_r_5r205_t23 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). Time Defined. (* constant 4669 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_ps1 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s s2), l_e_st_eq_landau_n_rt_cut))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s s2) => l_e_st_eq_landau_n_rt_rp_r_rpofp s (l_e_st_eq_landau_n_rt_rp_r_5r205_t33 s1 s2 p0 p1a p1b p2 case1 r a s i)))))))))))). Time Defined. (* constant 4670 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t34 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s s2), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_ps1 s1 s2 p0 p1a p1b p2 case1 r a s i)) s2))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s s2) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s2) s (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_ps1 s1 s2 p0 p1a p1b p2 case1 r a s i)) i (l_e_st_eq_landau_n_rt_rp_r_isprp1 s (l_e_st_eq_landau_n_rt_rp_r_5r205_t33 s1 s2 p0 p1a p1b p2 case1 r a s i))))))))))))). Time Defined. (* constant 4671 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t35 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s s2), l_e_st_nonempty l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s s2) => l_e_st_nonemptyi l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_ps1 s1 s2 p0 p1a p1b p2 case1 r a s i) (l_e_st_eq_landau_n_rt_rp_r_5r205_t27 s1 s2 p0 p1a p1b p2 case1 r a (l_e_st_eq_landau_n_rt_rp_r_5r205_ps1 s1 s2 p0 p1a p1b p2 case1 r a s i) (l_e_st_eq_landau_n_rt_rp_r_5r205_t34 s1 s2 p0 p1a p1b p2 case1 r a s i))))))))))))). Time Defined. (* constant 4672 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t36 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_nonempty l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_nonemptyapp l_e_st_eq_landau_n_rt_rp_r_real s2 p1b (l_e_st_nonempty l_e_st_eq_landau_n_rt_cut (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_in x s2) => l_e_st_eq_landau_n_rt_rp_r_5r205_t35 s1 s2 p0 p1a p1b p2 case1 r a x t))))))))))). Time Defined. (* constant 4673 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t37 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (j:l_e_st_eq_landau_n_rt_rp_in s0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)))))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (j:l_e_st_eq_landau_n_rt_rp_in s0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)) => p2 (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_5r205_t26 s1 s2 p0 p1a p1b p2 case1 r a r0 i) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_st_eq_landau_n_rt_rp_r_5r205_t28 s1 s2 p0 p1a p1b p2 case1 r a s0 j)))))))))))))). Time Defined. (* constant 4674 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t38 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (i:l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (j:l_e_st_eq_landau_n_rt_rp_in s0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_less r0 s0))))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (i:l_e_st_eq_landau_n_rt_rp_in r0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (j:l_e_st_eq_landau_n_rt_rp_in s0 (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_lessrpip r0 s0 (l_e_st_eq_landau_n_rt_rp_r_5r205_t37 s1 s2 p0 p1a p1b p2 case1 r a r0 i s0 j)))))))))))))). Time Defined. (* constant 4675 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_stc : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_cut))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_eq_landau_n_rt_rp_schnitt (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a) (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_5r205_t29 s1 s2 p0 p1a p1b p2 case1 r a x) (l_e_st_eq_landau_n_rt_rp_r_5r205_t31 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_t36 s1 s2 p0 p1a p1b p2 case1 r a) (fun (x:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (u:l_e_st_eq_landau_n_rt_rp_in y (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t38 s1 s2 p0 p1a p1b p2 case1 r a x t y u))))))))))))). Time Defined. (* constant 4676 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t39 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_less x (l_e_st_eq_landau_n_rt_rp_r_5r205_stc s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_eq_landau_n_rt_rp_satzp205a (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a) (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_5r205_t29 s1 s2 p0 p1a p1b p2 case1 r a x) (l_e_st_eq_landau_n_rt_rp_r_5r205_t31 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_t36 s1 s2 p0 p1a p1b p2 case1 r a) (fun (x:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (u:l_e_st_eq_landau_n_rt_rp_in y (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t38 s1 s2 p0 p1a p1b p2 case1 r a x t y u))))))))))))). Time Defined. (* constant 4677 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t40 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_all (fun (x:l_e_st_eq_landau_n_rt_cut) => (forall (t:l_e_st_eq_landau_n_rt_rp_more x (l_e_st_eq_landau_n_rt_rp_r_5r205_stc s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_eq_landau_n_rt_rp_satzp205b (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a) (fun (x:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_5r205_t29 s1 s2 p0 p1a p1b p2 case1 r a x) (l_e_st_eq_landau_n_rt_rp_r_5r205_t31 s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_t36 s1 s2 p0 p1a p1b p2 case1 r a) (fun (x:l_e_st_eq_landau_n_rt_cut) => (fun (t:l_e_st_eq_landau_n_rt_rp_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (y:l_e_st_eq_landau_n_rt_cut) => (fun (u:l_e_st_eq_landau_n_rt_rp_in y (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t38 s1 s2 p0 p1a p1b p2 case1 r a x t y u))))))))))))). Time Defined. (* constant 4678 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_stp : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_r_real))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_stc s1 s2 p0 p1a p1b p2 case1 r a)))))))))). Time Defined. (* constant 4679 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t41 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_e_st_eq_landau_n_rt_rp_r_posi (l_e_st_eq_landau_n_rt_rp_r_5r205_stc s1 s2 p0 p1a p1b p2 case1 r a)))))))))). Time Defined. (* constant 4680 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_ps2 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_cut)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_rpofp s p)))))))))))). Time Defined. (* constant 4681 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t42 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_less (l_e_st_eq_landau_n_rt_rp_r_5r205_ps2 s1 s2 p0 p1a p1b p2 case1 r a s l p) (l_e_st_eq_landau_n_rt_rp_r_5r205_stc s1 s2 p0 p1a p1b p2 case1 r a))))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_lessrpip (l_e_st_eq_landau_n_rt_rp_r_5r205_ps2 s1 s2 p0 p1a p1b p2 case1 r a s l p) (l_e_st_eq_landau_n_rt_rp_r_5r205_stc s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_isless1 s (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_ps2 s1 s2 p0 p1a p1b p2 case1 r a s l p)) (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_isprp1 s p) l))))))))))))). Time Defined. (* constant 4682 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t43 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_in (l_e_st_eq_landau_n_rt_rp_r_5r205_ps2 s1 s2 p0 p1a p1b p2 case1 r a s l p) (l_e_st_eq_landau_n_rt_rp_r_5r205_sc1 s1 s2 p0 p1a p1b p2 case1 r a))))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_5r205_t39 s1 s2 p0 p1a p1b p2 case1 r a (l_e_st_eq_landau_n_rt_rp_r_5r205_ps2 s1 s2 p0 p1a p1b p2 case1 r a s l p) (l_e_st_eq_landau_n_rt_rp_r_5r205_t42 s1 s2 p0 p1a p1b p2 case1 r a s l p))))))))))))). Time Defined. (* constant 4683 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t44 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos s), l_e_st_eq_landau_n_rt_rp_r_in s s1)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_ps2 s1 s2 p0 p1a p1b p2 case1 r a s l p)) s (l_e_st_eq_landau_n_rt_rp_r_5r205_t26 s1 s2 p0 p1a p1b p2 case1 r a (l_e_st_eq_landau_n_rt_rp_r_5r205_ps2 s1 s2 p0 p1a p1b p2 case1 r a s l p) (l_e_st_eq_landau_n_rt_rp_r_5r205_t43 s1 s2 p0 p1a p1b p2 case1 r a s l p)) (l_e_st_eq_landau_n_rt_rp_r_isprp2 s p))))))))))))). Time Defined. (* constant 4684 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t45 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_pos s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s s2), l_e_st_eq_landau_n_rt_rp_r_less r s))))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_pos s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s s2) => p2 r (l_e_st_eq_landau_n_rt_rp_r_5r205_t24 s1 s2 p0 p1a p1b p2 case1 r a) s i))))))))))))). Time Defined. (* constant 4685 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t46 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_pos s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in s s2), l_con))))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_pos s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in s s2) => n (l_e_st_eq_landau_n_rt_rp_r_lemma5 r s (l_e_st_eq_landau_n_rt_rp_r_5r205_t45 s1 s2 p0 p1a p1b p2 case1 r a s l n i) (l_e_st_eq_landau_n_rt_rp_r_5r205_t23 s1 s2 p0 p1a p1b p2 case1 r a))))))))))))))). Time Defined. (* constant 4686 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t47 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_pos s)), l_e_st_eq_landau_n_rt_rp_r_in s s1)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_pos s)) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_in s s1) (l_e_st_eq_landau_n_rt_rp_r_in s s2) (p0 s) (fun (t:l_e_st_eq_landau_n_rt_rp_r_in s s2) => l_e_st_eq_landau_n_rt_rp_r_5r205_t46 s1 s2 p0 p1a p1b p2 case1 r a s l n t))))))))))))). Time Defined. (* constant 4687 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t48 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_r_in s s1))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_pos s) (l_e_st_eq_landau_n_rt_rp_r_in s s1) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos s) => l_e_st_eq_landau_n_rt_rp_r_5r205_t44 s1 s2 p0 p1a p1b p2 case1 r a s l t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_pos s)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t47 s1 s2 p0 p1a p1b p2 case1 r a s l t)))))))))))). Time Defined. (* constant 4688 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t49 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_r_pos s))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_lemma5 (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a) s (l_e_st_eq_landau_n_rt_rp_r_lemma1 s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a) m) (l_e_st_eq_landau_n_rt_rp_r_5r205_t41 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). Time Defined. (* constant 4689 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_ps3 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_cut))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_rpofp s (l_e_st_eq_landau_n_rt_rp_r_5r205_t49 s1 s2 p0 p1a p1b p2 case1 r a s m)))))))))))). Time Defined. (* constant 4690 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t50 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_r_5r205_ps3 s1 s2 p0 p1a p1b p2 case1 r a s m) (l_e_st_eq_landau_n_rt_rp_r_5r205_stc s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_morerpip (l_e_st_eq_landau_n_rt_rp_r_5r205_ps3 s1 s2 p0 p1a p1b p2 case1 r a s m) (l_e_st_eq_landau_n_rt_rp_r_5r205_stc s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_ismore1 s (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_ps3 s1 s2 p0 p1a p1b p2 case1 r a s m)) (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_isprp1 s (l_e_st_eq_landau_n_rt_rp_r_5r205_t49 s1 s2 p0 p1a p1b p2 case1 r a s m)) m)))))))))))). Time Defined. (* constant 4691 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t51 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_in (l_e_st_eq_landau_n_rt_rp_r_5r205_ps3 s1 s2 p0 p1a p1b p2 case1 r a s m) (l_e_st_eq_landau_n_rt_rp_r_5r205_sc2 s1 s2 p0 p1a p1b p2 case1 r a)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t40 s1 s2 p0 p1a p1b p2 case1 r a (l_e_st_eq_landau_n_rt_rp_r_5r205_ps3 s1 s2 p0 p1a p1b p2 case1 r a s m) (l_e_st_eq_landau_n_rt_rp_r_5r205_t50 s1 s2 p0 p1a p1b p2 case1 r a s m)))))))))))). Time Defined. (* constant 4692 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t52 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)), l_e_st_eq_landau_n_rt_rp_r_in s s2))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s2) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_r_5r205_ps3 s1 s2 p0 p1a p1b p2 case1 r a s m)) s (l_e_st_eq_landau_n_rt_rp_r_5r205_t28 s1 s2 p0 p1a p1b p2 case1 r a (l_e_st_eq_landau_n_rt_rp_r_5r205_ps3 s1 s2 p0 p1a p1b p2 case1 r a s m) (l_e_st_eq_landau_n_rt_rp_r_5r205_t51 s1 s2 p0 p1a p1b p2 case1 r a s m)) (l_e_st_eq_landau_n_rt_rp_r_isprp2 s (l_e_st_eq_landau_n_rt_rp_r_5r205_t49 s1 s2 p0 p1a p1b p2 case1 r a s m))))))))))))). Time Defined. (* constant 4693 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t53 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_less x (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t48 s1 s2 p0 p1a p1b p2 case1 r a x t)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_more x (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t52 s1 s2 p0 p1a p1b p2 case1 r a x t))))))))))). Time Defined. (* constant 4694 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t54 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) (l_e_st_eq_landau_n_rt_rp_r_5r205_stp s1 s2 p0 p1a p1b p2 case1 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_t53 s1 s2 p0 p1a p1b p2 case1 r a)))))))))). Time Defined. (* constant 4695 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t55 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case1:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => l_someapp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)) case1 (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t54 s1 s2 p0 p1a p1b p2 case1 x t))))))))). Time Defined. (* constant 4696 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => l_e_st_setof l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 x) s1)))))))). Time Defined. (* constant 4697 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => l_e_st_setof l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 x) s2)))))))). Time Defined. (* constant 4698 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t56 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s1), l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s1) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 x) s1) r i))))))))). Time Defined. (* constant 4699 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t57 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s1))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)) => l_e_st_estie l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 x) s1) r i))))))))). Time Defined. (* constant 4700 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t58 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s2), l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s2) => l_e_st_estii l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 x) s2) r i))))))))). Time Defined. (* constant 4701 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t59 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s2))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) => l_e_st_estie l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 x) s2) r i))))))))). Time Defined. (* constant 4702 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t60 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_or (l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) (l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_comor (l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)) (l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) (l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s1) (l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s2) (l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)) (l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) (p0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s1) => l_e_st_eq_landau_n_rt_rp_r_5r205_t56 s1 s2 p0 p1a p1b p2 case2 r t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) s2) => l_e_st_eq_landau_n_rt_rp_r_5r205_t58 s1 s2 p0 p1a p1b p2 case2 r t)))))))))). Time Defined. (* constant 4703 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t61 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in r s2), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in r s2) => l_e_st_eq_landau_n_rt_rp_r_5r205_t58 s1 s2 p0 p1a p1b p2 case2 (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s2) r (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) i (l_e_st_eq_landau_n_rt_rp_r_satz177a r))))))))))). Time Defined. (* constant 4704 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t62 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in r s2), l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in r s2) => l_e_st_nonemptyi l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_t61 s1 s2 p0 p1a p1b p2 case2 r i)))))))))). Time Defined. (* constant 4705 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t63 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => l_e_st_nonemptyapp l_e_st_eq_landau_n_rt_rp_r_real s2 p1b (l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_in x s2) => l_e_st_eq_landau_n_rt_rp_r_5r205_t62 s1 s2 p0 p1a p1b p2 case2 x t))))))))). Time Defined. (* constant 4706 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t64 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in r s1), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in r s1) => l_e_st_eq_landau_n_rt_rp_r_5r205_t56 s1 s2 p0 p1a p1b p2 case2 (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s1) r (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) i (l_e_st_eq_landau_n_rt_rp_r_satz177a r))))))))))). Time Defined. (* constant 4707 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t65 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in r s1), l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in r s1) => l_e_st_nonemptyi l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_t64 s1 s2 p0 p1a p1b p2 case2 r i)))))))))). Time Defined. (* constant 4708 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t66 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => l_e_st_nonemptyapp l_e_st_eq_landau_n_rt_rp_r_real s1 p1a (l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_in x s1) => l_e_st_eq_landau_n_rt_rp_r_5r205_t65 s1 s2 p0 p1a p1b p2 case2 x t))))))))). Time Defined. (* constant 4709 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t67 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_m0 r)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)) => p2 (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_5r205_t57 s1 s2 p0 p1a p1b p2 case2 s j) (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_t59 s1 s2 p0 p1a p1b p2 case2 r i)))))))))))). Time Defined. (* constant 4710 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t68 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)), l_e_st_eq_landau_n_rt_rp_r_less r s))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_in r (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_in s (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)) => l_e_st_eq_landau_n_rt_rp_r_lemma1 s r (l_e_st_eq_landau_n_rt_rp_r_satz183d s r (l_e_st_eq_landau_n_rt_rp_r_5r205_t67 s1 s2 p0 p1a p1b p2 case2 r i s j))))))))))))). Time Defined. (* constant 4711 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t69 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)) => l_e_st_eq_landau_n_rt_rp_r_satz176c r (l_ande1 (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2) a)))))))))). Time Defined. (* constant 4712 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t70 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) s2))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s2) r (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_ande2 (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2) a) (l_e_st_eq_landau_n_rt_rp_r_satz177a r)))))))))). Time Defined. (* constant 4713 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t71 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)), l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) (l_e_st_eq_landau_n_rt_rp_r_5r205_t69 s1 s2 p0 p1a p1b p2 case2 r a) (l_e_st_eq_landau_n_rt_rp_r_5r205_t58 s1 s2 p0 p1a p1b p2 case2 (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_t70 s1 s2 p0 p1a p1b p2 case2 r a))))))))))). Time Defined. (* constant 4714 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t72 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2))) (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_t71 s1 s2 p0 p1a p1b p2 case2 r a)))))))))). Time Defined. (* constant 4715 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t73 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => l_someapp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)) case2 (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)))) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t72 s1 s2 p0 p1a p1b p2 case2 x t))))))))). Time Defined. (* constant 4716 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t74 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) x)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => l_e_st_eq_landau_n_rt_rp_r_5r205_t55 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_t60 s1 s2 p0 p1a p1b p2 case2 x) (l_e_st_eq_landau_n_rt_rp_r_5r205_t63 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_t66 s1 s2 p0 p1a p1b p2 case2) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_in x (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_in y (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t68 s1 s2 p0 p1a p1b p2 case2 x t y u)))) (l_e_st_eq_landau_n_rt_rp_r_5r205_t73 s1 s2 p0 p1a p1b p2 case2)))))))). Time Defined. (* constant 4717 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t75 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_m0 s) r))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_ismore2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_satz177 r) (l_e_st_eq_landau_n_rt_rp_r_satz183c s (l_e_st_eq_landau_n_rt_rp_r_m0 r) l)))))))))))). Time Defined. (* constant 4718 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t76 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) p (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_5r205_t75 s1 s2 p0 p1a p1b p2 case2 r p s l)))))))))))). Time Defined. (* constant 4719 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t77 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_in s s1))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less s (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s)) s (l_e_st_eq_landau_n_rt_rp_r_5r205_t57 s1 s2 p0 p1a p1b p2 case2 (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_5r205_t76 s1 s2 p0 p1a p1b p2 case2 r p s l)) (l_e_st_eq_landau_n_rt_rp_r_satz177 s)))))))))))). Time Defined. (* constant 4720 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t78 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_m0 s) r))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_isless2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) r (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_satz177 r) (l_e_st_eq_landau_n_rt_rp_r_satz183a s (l_e_st_eq_landau_n_rt_rp_r_m0 r) m)))))))))))). Time Defined. (* constant 4721 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t79 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_in (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) p (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_5r205_t78 s1 s2 p0 p1a p1b p2 case2 r p s m)))))))))))). Time Defined. (* constant 4722 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t80 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_in s s2))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more s (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_in x s2) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 s)) s (l_e_st_eq_landau_n_rt_rp_r_5r205_t59 s1 s2 p0 p1a p1b p2 case2 (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_5r205_t79 s1 s2 p0 p1a p1b p2 case2 r p s m)) (l_e_st_eq_landau_n_rt_rp_r_satz177 s)))))))))))). Time Defined. (* constant 4723 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t81 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r), l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_m0 r)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) => l_andi (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_less x (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t77 s1 s2 p0 p1a p1b p2 case2 r p x t)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_more x (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_5r205_t80 s1 s2 p0 p1a p1b p2 case2 r p x t))))))))))). Time Defined. (* constant 4724 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t82 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) r) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_t81 s1 s2 p0 p1a p1b p2 case2 r p)))))))))). Time Defined. (* constant 4725 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t83 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (case2:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => l_someapp l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) x) (l_e_st_eq_landau_n_rt_rp_r_5r205_t74 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 (l_e_st_eq_landau_n_rt_rp_r_5r205_sp2 s1 s2 p0 p1a p1b p2 case2) (l_e_st_eq_landau_n_rt_rp_r_5r205_sp1 s1 s2 p0 p1a p1b p2 case2) x) => l_e_st_eq_landau_n_rt_rp_r_5r205_t82 s1 s2 p0 p1a p1b p2 case2 x t))))))))). Time Defined. (* constant 4726 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t84 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), (forall (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0), l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => (fun (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0) => l_some_th4 l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)) notcase2 r)))))))))). Time Defined. (* constant 4727 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t85 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), (forall (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0), l_not (l_e_st_eq_landau_n_rt_rp_r_in r s2))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => (fun (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0) => l_and_th3 (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_in r s2) (l_e_st_eq_landau_n_rt_rp_r_5r205_t84 s1 s2 p0 p1a p1b p2 notcase1 notcase2 r l) (l_e_st_eq_landau_n_rt_rp_r_satz169d r l))))))))))). Time Defined. (* constant 4728 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t86 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), (forall (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_in r s1)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => (fun (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r l_e_st_eq_landau_n_rt_rp_r_0) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_in r s1) (l_e_st_eq_landau_n_rt_rp_r_in r s2) (p0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_t85 s1 s2 p0 p1a p1b p2 notcase1 notcase2 r l))))))))))). Time Defined. (* constant 4729 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t87 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), (forall (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0), l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1)))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => (fun (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) => l_some_th4 l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)) notcase1 r)))))))))). Time Defined. (* constant 4730 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t88 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), (forall (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0), l_not (l_e_st_eq_landau_n_rt_rp_r_in r s1))))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => (fun (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) => l_and_th3 (l_e_st_eq_landau_n_rt_rp_r_pos r) (l_e_st_eq_landau_n_rt_rp_r_in r s1) (l_e_st_eq_landau_n_rt_rp_r_5r205_t87 s1 s2 p0 p1a p1b p2 notcase1 notcase2 r m) (l_e_st_eq_landau_n_rt_rp_r_satz169b r m))))))))))). Time Defined. (* constant 4731 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t89 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), (forall (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_in r s2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => (fun (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) => l_ore2 (l_e_st_eq_landau_n_rt_rp_r_in r s1) (l_e_st_eq_landau_n_rt_rp_r_in r s2) (p0 r) (l_e_st_eq_landau_n_rt_rp_r_5r205_t88 s1 s2 p0 p1a p1b p2 notcase1 notcase2 r m))))))))))). Time Defined. (* constant 4732 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t90 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), (forall (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))), l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 l_e_st_eq_landau_n_rt_rp_r_0)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => (fun (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => l_andi (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 s1 s2 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 s1 s2 l_e_st_eq_landau_n_rt_rp_r_0) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_less x l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_5r205_t86 s1 s2 p0 p1a p1b p2 notcase1 notcase2 x t)) (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_more x l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_5r205_t89 s1 s2 p0 p1a p1b p2 notcase1 notcase2 x t)))))))))). Time Defined. (* constant 4733 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t91 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), (forall (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => (fun (notcase2:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_5r205_t90 s1 s2 p0 p1a p1b p2 notcase1 notcase2))))))))). Time Defined. (* constant 4734 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t92 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (notcase1:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => l_e_st_eq_landau_n_rt_rp_r_5r205_t83 s1 s2 p0 p1a p1b p2 t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_neg x) (l_e_st_eq_landau_n_rt_rp_r_in x s2)))) => l_e_st_eq_landau_n_rt_rp_r_5r205_t91 s1 s2 p0 p1a p1b p2 notcase1 t)))))))). Time Defined. (* constant 4735 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t93 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1))) => l_e_st_eq_landau_n_rt_rp_r_5r205_t55 s1 s2 p0 p1a p1b p2 t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_some (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_pos x) (l_e_st_eq_landau_n_rt_rp_r_in x s1)))) => l_e_st_eq_landau_n_rt_rp_r_5r205_t92 s1 s2 p0 p1a p1b p2 t))))))). Time Defined. (* constant 4736 *) Definition l_e_st_eq_landau_n_rt_rp_r_5r205_t94 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), l_e_st_eq_landau_n_rt_rp_r_one (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => l_e_onei l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) (l_e_st_eq_landau_n_rt_rp_r_5r205_t22 s1 s2 p0 p1a p1b p2) (l_e_st_eq_landau_n_rt_rp_r_5r205_t93 s1 s2 p0 p1a p1b p2))))))). Time Defined. (* constant 4737 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz205 : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), l_e_st_eq_landau_n_rt_rp_r_one (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_less y x), l_e_st_eq_landau_n_rt_rp_r_in y s1))) (l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_more y x), l_e_st_eq_landau_n_rt_rp_r_in y s2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => l_e_st_eq_landau_n_rt_rp_r_5r205_t94 s1 s2 p0 p1a p1b p2)))))). Time Defined. (* constant 4738 *) Definition l_e_st_eq_landau_n_rt_rp_r_dedekind : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), l_e_st_eq_landau_n_rt_rp_r_one (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_less y x), l_e_st_eq_landau_n_rt_rp_r_in y s1))) (l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_more y x), l_e_st_eq_landau_n_rt_rp_r_in y s2)))))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => l_e_st_eq_landau_n_rt_rp_r_satz205 s1 s2 p0 p1a p1b p2)))))). Time Defined. (* constant 4739 *) Definition l_e_st_eq_landau_n_rt_rp_r_schnitt : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), l_e_st_eq_landau_n_rt_rp_r_real)))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => l_e_ind l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) (l_e_st_eq_landau_n_rt_rp_r_satz205 s1 s2 p0 p1a p1b p2))))))). Time Defined. (* constant 4740 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz205a : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r (l_e_st_eq_landau_n_rt_rp_r_schnitt s1 s2 p0 p1a p1b p2)), l_e_st_eq_landau_n_rt_rp_r_in r s1)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r (l_e_st_eq_landau_n_rt_rp_r_schnitt s1 s2 p0 p1a p1b p2)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_schnitt s1 s2 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_schnitt s1 s2 p0 p1a p1b p2)) (l_e_oneax l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) (l_e_st_eq_landau_n_rt_rp_r_satz205 s1 s2 p0 p1a p1b p2)) r l)))))))). Time Defined. (* constant 4741 *) Definition l_e_st_eq_landau_n_rt_rp_r_satz205b : (forall (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real), (forall (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))), (forall (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1), (forall (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more r (l_e_st_eq_landau_n_rt_rp_r_schnitt s1 s2 p0 p1a p1b p2)), l_e_st_eq_landau_n_rt_rp_r_in r s2)))))))). exact (fun (s1:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s2:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p0:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_or (l_e_st_eq_landau_n_rt_rp_r_in x s1) (l_e_st_eq_landau_n_rt_rp_r_in x s2))) => (fun (p1a:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s1) => (fun (p1b:l_e_st_nonempty l_e_st_eq_landau_n_rt_rp_r_real s2) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_all (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_in x s1), l_e_st_eq_landau_n_rt_rp_r_all (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (forall (u:l_e_st_eq_landau_n_rt_rp_r_in y s2), l_e_st_eq_landau_n_rt_rp_r_less x y))))) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more r (l_e_st_eq_landau_n_rt_rp_r_schnitt s1 s2 p0 p1a p1b p2)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_5r205_prop1 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_schnitt s1 s2 p0 p1a p1b p2)) (l_e_st_eq_landau_n_rt_rp_r_5r205_prop2 s1 s2 (l_e_st_eq_landau_n_rt_rp_r_schnitt s1 s2 p0 p1a p1b p2)) (l_e_oneax l_e_st_eq_landau_n_rt_rp_r_real (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_5r205_prop3 s1 s2 x) (l_e_st_eq_landau_n_rt_rp_r_satz205 s1 s2 p0 p1a p1b p2)) r m)))))))). Time Defined. (* constant 4742 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_dr : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_dif)). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pdofrp r0)). Time Defined. (* constant 4743 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_ds : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_dif)). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pdofrp s0)). Time Defined. (* constant 4744 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t1 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_r_iva_dr r0 s0) (l_e_st_eq_landau_n_rt_rp_r_iva_ds r0 s0)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl r0 s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_lemmaivad1 r0 s0)). Time Defined. (* constant 4745 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemmaiva1 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_pl r0 s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_pl r0 s0)) (l_e_st_eq_landau_n_rt_rp_pd (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp s0)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl r0 s0)) (l_e_st_eq_landau_n_rt_rp_r_picp (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp s0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp s0))) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_pl r0 s0))) (l_e_st_eq_landau_n_rt_rp_r_iva_t1 r0 s0))). Time Defined. (* constant 4746 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t2 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_r_iva_dr r0 s0) (l_e_st_eq_landau_n_rt_rp_r_iva_ds r0 s0)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts r0 s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_lemmaivad2 r0 s0)). Time Defined. (* constant 4747 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemmaiva2 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_ts r0 s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_ts r0 s0)) (l_e_st_eq_landau_n_rt_rp_td (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp s0)) (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts r0 s0)) (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_st_eq_landau_n_rt_rp_pdofrp r0) (l_e_st_eq_landau_n_rt_rp_pdofrp s0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp r0)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp s0))) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_pdofrp (l_e_st_eq_landau_n_rt_rp_ts r0 s0))) (l_e_st_eq_landau_n_rt_rp_r_iva_t2 r0 s0))). Time Defined. (* constant 4748 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t3 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)), l_e_st_eq_landau_n_rt_rp_mored (l_e_st_eq_landau_n_rt_rp_r_iva_dr r0 s0) (l_e_st_eq_landau_n_rt_rp_r_iva_ds r0 s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_r_moreex (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) (l_e_st_eq_landau_n_rt_rp_r_iva_dr r0 s0) (l_e_st_eq_landau_n_rt_rp_r_iva_ds r0 s0) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_r_iva_dr r0 s0)) (l_e_st_eq_landau_n_rt_rp_r_innclass (l_e_st_eq_landau_n_rt_rp_r_iva_ds r0 s0)) m))). Time Defined. (* constant 4749 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemmaiva3 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)), l_e_st_eq_landau_n_rt_rp_more r0 s0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_lemmaivad3 r0 s0 (l_e_st_eq_landau_n_rt_rp_r_iva_t3 r0 s0 m)))). Time Defined. (* constant 4750 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t4 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r0 s0), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)), l_e_st_eq_landau_n_rt_rp_less r0 s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r0 s0) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_satz121 s0 r0 (l_e_st_eq_landau_n_rt_rp_r_lemmaiva3 s0 r0 (l_e_st_eq_landau_n_rt_rp_r_lemma2 (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0) l)))))). Time Defined. (* constant 4751 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t5 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r0 s0), l_not (l_e_st_eq_landau_n_rt_rp_less r0 s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r0 s0) => l_ec3e23 (l_e_st_eq_landau_n_rt_rp_is r0 s0) (l_e_st_eq_landau_n_rt_rp_more r0 s0) (l_e_st_eq_landau_n_rt_rp_less r0 s0) (l_e_st_eq_landau_n_rt_rp_satz123b r0 s0) m))). Time Defined. (* constant 4752 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t6 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r0 s0), l_not (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0))))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r0 s0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_less r0 s0) (l_e_st_eq_landau_n_rt_rp_r_iva_t5 r0 s0 m) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_r_iva_t4 r0 s0 m t)))). Time Defined. (* constant 4753 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t7 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r0 s0), l_e_st_eq_landau_n_rt_rp_nis r0 s0))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r0 s0) => l_ec3e21 (l_e_st_eq_landau_n_rt_rp_is r0 s0) (l_e_st_eq_landau_n_rt_rp_more r0 s0) (l_e_st_eq_landau_n_rt_rp_less r0 s0) (l_e_st_eq_landau_n_rt_rp_satz123b r0 s0) m))). Time Defined. (* constant 4754 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t8 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r0 s0), l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r0 s0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_is r0 s0) (l_e_st_eq_landau_n_rt_rp_r_iva_t7 r0 s0 m) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) => l_e_st_eq_landau_n_rt_rp_r_isrpip r0 s0 t)))). Time Defined. (* constant 4755 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemmaiva4 : (forall (r0:l_e_st_eq_landau_n_rt_cut), (forall (s0:l_e_st_eq_landau_n_rt_cut), (forall (m:l_e_st_eq_landau_n_rt_rp_more r0 s0), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)))). exact (fun (r0:l_e_st_eq_landau_n_rt_cut) => (fun (s0:l_e_st_eq_landau_n_rt_cut) => (fun (m:l_e_st_eq_landau_n_rt_rp_more r0 s0) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_r_satz167a (l_e_st_eq_landau_n_rt_rp_r_pofrp r0) (l_e_st_eq_landau_n_rt_rp_r_pofrp s0)) (l_e_st_eq_landau_n_rt_rp_r_iva_t6 r0 s0 m) (l_e_st_eq_landau_n_rt_rp_r_iva_t8 r0 s0 m)))). Time Defined. (* constant 4756 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t9 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) => l_e_st_eq_landau_n_rt_rp_r_lemmaiva3 (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y) m))). Time Defined. (* constant 4757 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t10 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) => l_e_st_eq_landau_n_rt_rp_satz154d (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y) (l_e_st_eq_landau_n_rt_rp_r_iva_t9 x y m)))). Time Defined. (* constant 4758 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t11 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) => l_e_st_eq_landau_n_rt_moree (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y) (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_iva_t10 x y m)))). Time Defined. (* constant 4759 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemmaiva5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)), l_e_st_eq_landau_n_more x y))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) => l_e_st_eq_landau_n_satz111a x y (l_e_st_eq_landau_n_rt_rp_r_iva_t11 x y m)))). Time Defined. (* constant 4760 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_moref (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_satz111d x y m))). Time Defined. (* constant 4761 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t13 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_rt_more (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_rt_morei (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y) (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_inclass (l_e_st_eq_landau_n_fr y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_iva_t12 x y m)))). Time Defined. (* constant 4762 *) Definition l_e_st_eq_landau_n_rt_rp_r_iva_t14 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_rt_rp_more (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_rt_rp_satz154a (l_e_st_eq_landau_n_rt_rtofn x) (l_e_st_eq_landau_n_rt_rtofn y) (l_e_st_eq_landau_n_rt_rp_r_iva_t13 x y m)))). Time Defined. (* constant 4763 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemmaiva6 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_more x y), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_more x y) => l_e_st_eq_landau_n_rt_rp_r_lemmaiva4 (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y) (l_e_st_eq_landau_n_rt_rp_r_iva_t14 x y m)))). Time Defined. (* constant 4764 *) Definition l_e_st_eq_landau_n_rt_rp_r_int_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_absd a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_intabsd a0 (l_e_st_eq_landau_n_rt_rp_r_intrlex r a0 a0ir i))))). Time Defined. (* constant 4765 *) Definition l_e_st_eq_landau_n_rt_rp_r_int_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_abs r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_r_intrlin (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_absd a0) (l_e_st_eq_landau_n_rt_rp_r_aica r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_int_t1 r a0 a0ir i))))). Time Defined. (* constant 4766 *) Definition l_e_st_eq_landau_n_rt_rp_r_intabs : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_abs r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_abs r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_int_t2 r x t i)))). Time Defined. (* constant 4767 *) Definition l_e_st_eq_landau_n_rt_rp_r_int_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_m0d a0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_intm0d a0 (l_e_st_eq_landau_n_rt_rp_r_intrlex r a0 a0ir i))))). Time Defined. (* constant 4768 *) Definition l_e_st_eq_landau_n_rt_rp_r_int_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_m0 r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a0ir:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_r_intrlin (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_m0d a0) (l_e_st_eq_landau_n_rt_rp_r_micm0 r a0 a0ir) (l_e_st_eq_landau_n_rt_rp_r_int_t3 r a0 a0ir i))))). Time Defined. (* constant 4769 *) Definition l_e_st_eq_landau_n_rt_rp_r_intm0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_m0 r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => l_e_st_eq_landau_n_rt_rp_r_realapp1 r (l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => l_e_st_eq_landau_n_rt_rp_r_int_t4 r x t i)))). Time Defined. (* constant 4770 *) Definition l_e_st_eq_landau_n_rt_rp_r_int_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_pd a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => l_e_st_eq_landau_n_rt_rp_intpd a1 b1 (l_e_st_eq_landau_n_rt_rp_r_intrlex r a1 a1ir i) (l_e_st_eq_landau_n_rt_rp_r_intrlex s b1 b1is j))))))))). Time Defined. (* constant 4771 *) Definition l_e_st_eq_landau_n_rt_rp_r_int_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_pl r s))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => l_e_st_eq_landau_n_rt_rp_r_intrlin (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_pd a1 b1) (l_e_st_eq_landau_n_rt_rp_r_picp r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_int_t5 r s a1 b1 a1ir b1is i j))))))))). Time Defined. (* constant 4772 *) Definition l_e_st_eq_landau_n_rt_rp_r_intpl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_pl r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_pl r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_int_t6 r s x y t u i j)))))))). Time Defined. (* constant 4773 *) Definition l_e_st_eq_landau_n_rt_rp_r_intmn : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_mn r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => l_e_st_eq_landau_n_rt_rp_r_intpl r i (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_intm0 s j))))). Time Defined. (* constant 4774 *) Definition l_e_st_eq_landau_n_rt_rp_r_int_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), l_e_st_eq_landau_n_rt_rp_intd (l_e_st_eq_landau_n_rt_rp_td a1 b1))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => l_e_st_eq_landau_n_rt_rp_inttd a1 b1 (l_e_st_eq_landau_n_rt_rp_r_intrlex r a1 a1ir i) (l_e_st_eq_landau_n_rt_rp_r_intrlex s b1 b1is j))))))))). Time Defined. (* constant 4775 *) Definition l_e_st_eq_landau_n_rt_rp_r_int_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a1:l_e_st_eq_landau_n_rt_rp_dif), (forall (b1:l_e_st_eq_landau_n_rt_rp_dif), (forall (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_ts r s))))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (b1:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a1ir:l_e_st_eq_landau_n_rt_rp_r_inn a1 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b1is:l_e_st_eq_landau_n_rt_rp_r_inn b1 (l_e_st_eq_landau_n_rt_rp_r_class s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => l_e_st_eq_landau_n_rt_rp_r_intrlin (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_td a1 b1) (l_e_st_eq_landau_n_rt_rp_r_tict r s a1 b1 a1ir b1is) (l_e_st_eq_landau_n_rt_rp_r_int_t7 r s a1 b1 a1ir b1is i j))))))))). Time Defined. (* constant 4776 *) Definition l_e_st_eq_landau_n_rt_rp_r_intts : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_ts r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => l_e_st_eq_landau_n_rt_rp_r_realapp2 r s (l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_ts r s)) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_int_t8 r s x y t u i j)))))))). Time Defined. (* constant 4777 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr24_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n))). Time Defined. (* constant 4778 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr24_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_not (l_e_st_eq_landau_n_rt_rp_r_more l_e_st_eq_landau_n_rt_rp_r_1rl r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_more l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_more l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n)) (l_e_st_eq_landau_n_satz10d l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n) (l_e_st_eq_landau_n_rt_rp_r_ivr24_t1 r n)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_more l_e_st_eq_landau_n_rt_rp_r_1rl r) => l_e_st_eq_landau_n_rt_rp_r_lemmaiva5 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n) (l_e_st_eq_landau_n_rt_rp_r_ismore2 r (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ntofrl r n)) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 r n) t)))). Time Defined. (* constant 4779 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr24 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_natrl r), l_e_st_eq_landau_n_rt_rp_r_lessis l_e_st_eq_landau_n_rt_rp_r_1rl r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_natrl r) => l_e_st_eq_landau_n_rt_rp_r_satz167e l_e_st_eq_landau_n_rt_rp_r_1rl r (l_e_st_eq_landau_n_rt_rp_r_ivr24_t2 r n))). Time Defined. (* constant 4780 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr25_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn s r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_satz182d s r (l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l)))))). Time Defined. (* constant 4781 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr25_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_mn s r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_intmn s j r i))))). Time Defined. (* constant 4782 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr25_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_natrl (l_e_st_eq_landau_n_rt_rp_r_mn s r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_posintnatrl (l_e_st_eq_landau_n_rt_rp_r_mn s r) (l_e_st_eq_landau_n_rt_rp_r_ivr25_t1 r i s j l) (l_e_st_eq_landau_n_rt_rp_r_ivr25_t2 r i s j l)))))). Time Defined. (* constant 4783 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr25_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_lessis l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_mn s r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_satzr24 (l_e_st_eq_landau_n_rt_rp_r_mn s r) (l_e_st_eq_landau_n_rt_rp_r_ivr25_t3 r i s j l)))))). Time Defined. (* constant 4784 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr25_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn s r) r)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_less l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_mn s r)) (l_e_st_eq_landau_n_rt_rp_r_is l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_mn s r)) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn s r) r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn s r) r)) (l_e_st_eq_landau_n_rt_rp_r_ivr25_t4 r i s j l) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_mn s r)) => l_e_st_eq_landau_n_rt_rp_r_satz188f l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_mn s r) r t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_mn s r)) => l_e_st_eq_landau_n_rt_rp_r_ispl1 l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_mn s r) r t)))))). Time Defined. (* constant 4785 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr25 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_intrl r), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (j:l_e_st_eq_landau_n_rt_rp_r_intrl s), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_1rl) s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_intrl r) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_intrl s) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_islessis12 (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn s r) r) s (l_e_st_eq_landau_n_rt_rp_r_compl l_e_st_eq_landau_n_rt_rp_r_1rl r) (l_e_st_eq_landau_n_rt_rp_r_plmn s r) (l_e_st_eq_landau_n_rt_rp_r_ivr25_t5 r i s j l)))))). Time Defined. (* constant 4786 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr155_t1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_lemmaiva1 (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y))). Time Defined. (* constant 4787 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr155_t2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_isrpep (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)) (l_e_st_eq_landau_n_rt_rp_satz155e x y))). Time Defined. (* constant 4788 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr155a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_pl (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y))) (l_e_st_eq_landau_n_rt_rp_r_ivr155_t2 x y) (l_e_st_eq_landau_n_rt_rp_r_ivr155_t1 x y))). Time Defined. (* constant 4789 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr155b : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) (l_e_st_eq_landau_n_rt_rp_r_satzr155a x y))). Time Defined. (* constant 4790 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr155_t3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_lemmaiva2 (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y))). Time Defined. (* constant 4791 *) Definition l_e_st_eq_landau_n_rt_rp_r_ivr155_t4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_isrpep (l_e_st_eq_landau_n_rt_rp_rpofnt (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y)) (l_e_st_eq_landau_n_rt_rp_satz155f x y))). Time Defined. (* constant 4792 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr155c : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) (l_e_st_eq_landau_n_rt_rp_r_pofrp (l_e_st_eq_landau_n_rt_rp_ts (l_e_st_eq_landau_n_rt_rp_rpofnt x) (l_e_st_eq_landau_n_rt_rp_rpofnt y))) (l_e_st_eq_landau_n_rt_rp_r_ivr155_t4 x y) (l_e_st_eq_landau_n_rt_rp_r_ivr155_t3 x y))). Time Defined. (* constant 4793 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr155d : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_ts x y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) (l_e_st_eq_landau_n_rt_rp_r_rlofnt y)) (l_e_st_eq_landau_n_rt_rp_r_satzr155c x y))). Time Defined. (* constant 4794 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t1 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)), (forall (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b0:l_e_st_eq_landau_n_rt_rp_dif), (forall (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_not (l_e_st_eq_landau_n_rt_rp_negd a0)))))))))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)) => (fun (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd a0) (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_ande1 (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t) a) (fun (u:l_e_st_eq_landau_n_rt_rp_negd a0) => l_e_st_eq_landau_n_rt_rp_r_negin r a0 air u)))))))))))). Time Defined. (* constant 4795 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t2 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)), (forall (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b0:l_e_st_eq_landau_n_rt_rp_dif), (forall (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_not (l_e_st_eq_landau_n_rt_rp_negd b0)))))))))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)) => (fun (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd b0) (l_e_st_eq_landau_n_rt_rp_r_neg s) (l_ande1 (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t) b) (fun (u:l_e_st_eq_landau_n_rt_rp_negd b0) => l_e_st_eq_landau_n_rt_rp_r_negin s b0 bis u)))))))))))). Time Defined. (* constant 4796 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t3 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)), (forall (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b0:l_e_st_eq_landau_n_rt_rp_dif), (forall (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0))))))))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)) => (fun (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_ts r r) t (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0 (l_e_st_eq_landau_n_rt_rp_r_tict r r a0 a0 air air) cit (l_ande2 (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t) a)))))))))))). Time Defined. (* constant 4797 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t4 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)), (forall (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b0:l_e_st_eq_landau_n_rt_rp_dif), (forall (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td b0 b0) c0))))))))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)) => (fun (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isex (l_e_st_eq_landau_n_rt_rp_r_ts s s) t (l_e_st_eq_landau_n_rt_rp_td b0 b0) c0 (l_e_st_eq_landau_n_rt_rp_r_tict s s b0 b0 bis bis) cit (l_ande2 (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t) b)))))))))))). Time Defined. (* constant 4798 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t5 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)), (forall (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)), (forall (b0:l_e_st_eq_landau_n_rt_rp_dif), (forall (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)), l_e_st_eq_landau_n_rt_rp_r_is r s))))))))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)) => (fun (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (air:l_e_st_eq_landau_n_rt_rp_r_inn a0 (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (b0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (bis:l_e_st_eq_landau_n_rt_rp_r_inn b0 (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_isin r s a0 b0 air bis (l_e_st_eq_landau_n_rt_rp_satzd161b c0 a0 b0 (l_e_st_eq_landau_n_rt_rp_r_7r161_t1 t r s a b c0 cit a0 air b0 bis) (l_e_st_eq_landau_n_rt_rp_r_7r161_t2 t r s a b c0 cit a0 air b0 bis) (l_e_st_eq_landau_n_rt_rp_r_7r161_t3 t r s a b c0 cit a0 air b0 bis) (l_e_st_eq_landau_n_rt_rp_r_7r161_t4 t r s a b c0 cit a0 air b0 bis))))))))))))). Time Defined. (* constant 4799 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr161b : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)), (forall (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)), l_e_st_eq_landau_n_rt_rp_r_is r s))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) t)) => (fun (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) t)) => l_e_st_eq_landau_n_rt_rp_r_realapp3 t r s (l_e_st_eq_landau_n_rt_rp_r_is r s) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (y:l_e_st_eq_landau_n_rt_rp_dif) => (fun (z:l_e_st_eq_landau_n_rt_rp_dif) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn y (l_e_st_eq_landau_n_rt_rp_r_class r)) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_inn z (l_e_st_eq_landau_n_rt_rp_r_class s)) => l_e_st_eq_landau_n_rt_rp_r_7r161_t5 t r s a b x u y v z w))))))))))). Time Defined. (* constant 4800 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t6 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), l_not (l_e_st_eq_landau_n_rt_rp_negd c0))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_negd c0) (l_e_st_eq_landau_n_rt_rp_r_neg t) n (fun (u:l_e_st_eq_landau_n_rt_rp_negd c0) => l_e_st_eq_landau_n_rt_rp_r_negin t c0 cit u))))). Time Defined. (* constant 4801 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_ar : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)), l_e_st_eq_landau_n_rt_rp_r_real)))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)) => l_e_st_eq_landau_n_rt_rp_r_realof a0)))))). Time Defined. (* constant 4802 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t7 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)), l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a)))))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a)) (l_e_st_eq_landau_n_rt_rp_negd a0) (l_ande1 (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0) a) (fun (u:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a)) => l_e_st_eq_landau_n_rt_rp_r_negex (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a) a0 (l_e_st_eq_landau_n_rt_rp_r_innclass a0) u))))))). Time Defined. (* constant 4803 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t8 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a) (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a)) t)))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)) => l_e_st_eq_landau_n_rt_rp_r_isin (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a) (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a)) t (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0 (l_e_st_eq_landau_n_rt_rp_r_tict (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a) (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a) a0 a0 (l_e_st_eq_landau_n_rt_rp_r_innclass a0) (l_e_st_eq_landau_n_rt_rp_r_innclass a0)) cit (l_ande2 (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0) a))))))). Time Defined. (* constant 4804 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t9 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)), l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a))) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a) (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a)) t))))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)) => l_andi (l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a))) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a) (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a)) t) (l_e_st_eq_landau_n_rt_rp_r_7r161_t7 t n c0 cit a0 a) (l_e_st_eq_landau_n_rt_rp_r_7r161_t8 t n c0 cit a0 a))))))). Time Defined. (* constant 4805 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t10 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), (forall (a0:l_e_st_eq_landau_n_rt_rp_dif), (forall (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)), l_e_st_eq_landau_n_rt_rp_r_some (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)))))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => (fun (a0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd a0)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td a0 a0) c0)) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)) (l_e_st_eq_landau_n_rt_rp_r_7r161_ar t n c0 cit a0 a) (l_e_st_eq_landau_n_rt_rp_r_7r161_t9 t n c0 cit a0 a))))))). Time Defined. (* constant 4806 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t11 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (c0:l_e_st_eq_landau_n_rt_rp_dif), (forall (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)), l_e_st_eq_landau_n_rt_rp_r_some (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (c0:l_e_st_eq_landau_n_rt_rp_dif) => (fun (cit:l_e_st_eq_landau_n_rt_rp_r_inn c0 (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_someapp l_e_st_eq_landau_n_rt_rp_dif (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd x)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td x x) c0)) (l_e_st_eq_landau_n_rt_rp_satzd161a c0 (l_e_st_eq_landau_n_rt_rp_r_7r161_t6 t n c0 cit)) (l_e_st_eq_landau_n_rt_rp_r_some (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (v:l_and (l_not (l_e_st_eq_landau_n_rt_rp_negd x)) (l_e_st_eq_landau_n_rt_rp_eq (l_e_st_eq_landau_n_rt_rp_td x x) c0)) => l_e_st_eq_landau_n_rt_rp_r_7r161_t10 t n c0 cit x v)))))). Time Defined. (* constant 4807 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr161a : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), l_e_st_eq_landau_n_rt_rp_r_some (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => l_e_st_eq_landau_n_rt_rp_r_realapp1 t (l_e_st_eq_landau_n_rt_rp_r_some (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t))) (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_inn x (l_e_st_eq_landau_n_rt_rp_r_class t)) => l_e_st_eq_landau_n_rt_rp_r_7r161_t11 t n x v)))). Time Defined. (* constant 4808 *) Definition l_e_st_eq_landau_n_rt_rp_r_satzr161 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), l_e_st_eq_landau_n_rt_rp_r_one (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => l_e_onei l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)) (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)) => (fun (b:l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg v)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts v v) t)) => l_e_st_eq_landau_n_rt_rp_r_satzr161b t u v a b)))) (l_e_st_eq_landau_n_rt_rp_r_satzr161a t n))). Time Defined. (* constant 4809 *) Definition l_e_st_eq_landau_n_rt_rp_r_sqrt : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => l_e_ind l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)) (l_e_st_eq_landau_n_rt_rp_r_satzr161 t n))). Time Defined. (* constant 4810 *) Definition l_e_st_eq_landau_n_rt_rp_r_7r161_t12 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_sqrt t n))) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_sqrt t n) (l_e_st_eq_landau_n_rt_rp_r_sqrt t n)) t))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => l_e_oneax l_e_st_eq_landau_n_rt_rp_r_real (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg u)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts u u) t)) (l_e_st_eq_landau_n_rt_rp_r_satzr161 t n))). Time Defined. (* constant 4811 *) Definition l_e_st_eq_landau_n_rt_rp_r_thsqrt1a : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_sqrt t n)))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => l_ande1 (l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_sqrt t n))) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_sqrt t n) (l_e_st_eq_landau_n_rt_rp_r_sqrt t n)) t) (l_e_st_eq_landau_n_rt_rp_r_7r161_t12 t n))). Time Defined. (* constant 4812 *) Definition l_e_st_eq_landau_n_rt_rp_r_thsqrt1b : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_sqrt t n) (l_e_st_eq_landau_n_rt_rp_r_sqrt t n)) t)). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => l_ande2 (l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_sqrt t n))) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_sqrt t n) (l_e_st_eq_landau_n_rt_rp_r_sqrt t n)) t) (l_e_st_eq_landau_n_rt_rp_r_7r161_t12 t n))). Time Defined. (* constant 4813 *) Definition l_e_st_eq_landau_n_rt_rp_r_thsqrt2 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg x)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts x x) t), l_e_st_eq_landau_n_rt_rp_r_is x (l_e_st_eq_landau_n_rt_rp_r_sqrt t n)))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg x)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts x x) t) => l_e_st_eq_landau_n_rt_rp_r_satzr161b t x (l_e_st_eq_landau_n_rt_rp_r_sqrt t n) (l_andi (l_not (l_e_st_eq_landau_n_rt_rp_r_neg x)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts x x) t) o i) (l_e_st_eq_landau_n_rt_rp_r_7r161_t12 t n)))))). Time Defined. (* constant 4814 *) Definition l_e_st_eq_landau_n_rt_rp_r_thsqrt3 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)), (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg x)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is t (l_e_st_eq_landau_n_rt_rp_r_ts x x)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_sqrt t n) x))))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) => (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg x)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is t (l_e_st_eq_landau_n_rt_rp_r_ts x x)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real x (l_e_st_eq_landau_n_rt_rp_r_sqrt t n) (l_e_st_eq_landau_n_rt_rp_r_thsqrt2 t n x o (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real t (l_e_st_eq_landau_n_rt_rp_r_ts x x) i))))))). Time Defined. (* constant 4815 *) Definition l_e_st_eq_landau_n_rt_rp_r_issqrt : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) (l_e_st_eq_landau_n_rt_rp_r_sqrt s o)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_thsqrt2 s o (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) (l_e_st_eq_landau_n_rt_rp_r_thsqrt1a r n) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) (l_e_st_eq_landau_n_rt_rp_r_sqrt r n)) r s (l_e_st_eq_landau_n_rt_rp_r_thsqrt1b r n) i)))))). Time Defined. (* constant 4816 *) Definition l_e_st_eq_landau_n_rt_rp_r_sqrt0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_thsqrt3 r n l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_0notn l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 i (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))))). Time Defined. (* constant 4817 *) Definition l_e_st_eq_landau_n_rt_rp_r_sqrt_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (o:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) o (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) (l_e_st_eq_landau_n_rt_rp_r_sqrt r n)) (l_e_st_eq_landau_n_rt_rp_r_thsqrt1b r n) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) t))))). Time Defined. (* constant 4818 *) Definition l_e_st_eq_landau_n_rt_rp_r_sqrtnot0 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (o:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_sqrt r n)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_or3e2 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_sqrt r n) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_sqrt r n)) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_sqrt r n)) (l_e_st_eq_landau_n_rt_rp_r_axrlo (l_e_st_eq_landau_n_rt_rp_r_sqrt r n)) (l_e_st_eq_landau_n_rt_rp_r_thsqrt1a r n) (l_e_st_eq_landau_n_rt_rp_r_sqrt_t1 r n o)))). Time Defined. (* constant 4819 *) Definition l_e_st_eq_landau_n_rt_rp_r_v0_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis t l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ov s t n))) (l_e_st_eq_landau_n_rt_rp_r_ts r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis t l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ov s t n))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ov s t n)) t) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ov s t n) t)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_comts t (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ov s t n))) (l_e_st_eq_landau_n_rt_rp_r_assts1 r (l_e_st_eq_landau_n_rt_rp_r_ov s t n) t) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ov s t n) t) s r (l_e_st_eq_landau_n_rt_rp_r_satz204e s t n)))))). Time Defined. (* constant 4820 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis t l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ov s t n)) (l_e_st_eq_landau_n_rt_rp_r_ov (l_e_st_eq_landau_n_rt_rp_r_ts r s) t n))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis t l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz204g (l_e_st_eq_landau_n_rt_rp_r_ts r s) t (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_ov s t n)) n (l_e_st_eq_landau_n_rt_rp_r_v0_t1 r s t n))))). Time Defined. (* constant 4821 *) Definition l_e_st_eq_landau_n_rt_rp_r_v0_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis t l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ov r t n) (l_e_st_eq_landau_n_rt_rp_r_ov s t n))) (l_e_st_eq_landau_n_rt_rp_r_pl r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis t l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ov r t n) (l_e_st_eq_landau_n_rt_rp_r_ov s t n))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_ov r t n)) (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_ov s t n))) (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_disttp2 t (l_e_st_eq_landau_n_rt_rp_r_ov r t n) (l_e_st_eq_landau_n_rt_rp_r_ov s t n)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_ov r t n)) r (l_e_st_eq_landau_n_rt_rp_r_ts t (l_e_st_eq_landau_n_rt_rp_r_ov s t n)) s (l_e_st_eq_landau_n_rt_rp_r_satz204c r t n) (l_e_st_eq_landau_n_rt_rp_r_satz204c s t n)))))). Time Defined. (* constant 4822 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis t l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ov r t n) (l_e_st_eq_landau_n_rt_rp_r_ov s t n)) (l_e_st_eq_landau_n_rt_rp_r_ov (l_e_st_eq_landau_n_rt_rp_r_pl r s) t n))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis t l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz204g (l_e_st_eq_landau_n_rt_rp_r_pl r s) t (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ov r t n) (l_e_st_eq_landau_n_rt_rp_r_ov s t n)) n (l_e_st_eq_landau_n_rt_rp_r_v0_t2 r s t n))))). Time Defined. (* constant 4823 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ov r r n) l_e_st_eq_landau_n_rt_rp_r_1rl)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz204b r r n (l_e_st_eq_landau_n_rt_rp_r_ov r r n) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_satz204c r r n) (l_e_st_eq_landau_n_rt_rp_r_satz195 r))). Time Defined. (* constant 4824 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ov l_e_st_eq_landau_n_rt_rp_r_0 r n) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_ore2 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ov l_e_st_eq_landau_n_rt_rp_r_0 r n) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz192c r (l_e_st_eq_landau_n_rt_rp_r_ov l_e_st_eq_landau_n_rt_rp_r_0 r n) (l_e_st_eq_landau_n_rt_rp_r_satz204c l_e_st_eq_landau_n_rt_rp_r_0 r n)) n)). Time Defined. (* constant 4825 *) Definition l_e_st_eq_landau_n_rt_rp_r_v0_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 r)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_con))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_pnotn (l_e_st_eq_landau_n_rt_rp_r_m0 r) p (l_e_st_eq_landau_n_rt_rp_r_isneg r (l_e_st_eq_landau_n_rt_rp_r_m0 r) i (l_e_st_eq_landau_n_rt_rp_r_satz176f r p))))). Time Defined. (* constant 4826 *) Definition l_e_st_eq_landau_n_rt_rp_r_v0_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 r)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_con))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_nnotp (l_e_st_eq_landau_n_rt_rp_r_m0 r) n (l_e_st_eq_landau_n_rt_rp_r_ispos r (l_e_st_eq_landau_n_rt_rp_r_m0 r) i (l_e_st_eq_landau_n_rt_rp_r_satz176d r n))))). Time Defined. (* constant 4827 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 r)), l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_satz176e r (l_or3e1 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_axrlo (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_v0_t3 r i t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_m0 r)) => l_e_st_eq_landau_n_rt_rp_r_v0_t4 r i t)))). Time Defined. (* constant 4828 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_satz167f (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0 (l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts r r)) (l_e_st_eq_landau_n_rt_rp_r_nnegsq r) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz169d (l_e_st_eq_landau_n_rt_rp_r_ts r r) t))). Time Defined. (* constant 4829 *) Definition l_e_st_eq_landau_n_rt_rp_r_lemma12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_rapp r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs r))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_satz196a r r t t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs r)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 r r t) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_abs0 r t))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_satz196b r r t t)). Time Defined. (* constant 4830 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_0)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => l_e_st_eq_landau_n_rt_rp_r_satz190a x x l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_moreisi2 x x (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real x)) (l_e_st_eq_landau_n_rt_rp_r_satz169a l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_natpos l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_natrl1))))))). Time Defined. (* constant 4831 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) x))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => l_e_st_eq_landau_n_rt_rp_r_ismore2 (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_0) x (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl02 x l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_shift_t1 x ix y iy ly)))))). Time Defined. (* constant 4832 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => l_e_st_eq_landau_n_rt_rp_r_satz172d (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) x y (l_e_st_eq_landau_n_rt_rp_r_shift_t2 x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_satz168b y x ly)))))). Time Defined. (* constant 4833 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => l_e_st_eq_landau_n_rt_rp_r_satz182d (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y (l_e_st_eq_landau_n_rt_rp_r_shift_t3 x ix y iy ly)))))). Time Defined. (* constant 4834 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => l_e_st_eq_landau_n_rt_rp_r_intmn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1) y iy))))). Time Defined. (* constant 4835 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), l_e_st_eq_landau_n_rt_rp_r_natrl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => l_e_st_eq_landau_n_rt_rp_r_posintnatrl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t4 x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_shift_t5 x ix y iy ly)))))). Time Defined. (* constant 4836 *) Definition l_e_st_eq_landau_n_rt_rp_r_shiftl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), l_e_st_eq_landau_n_nat))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 x ix y iy ly)))))). Time Defined. (* constant 4837 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_n1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_nat)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_inn (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) n)))))). Time Defined. (* constant 4838 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) n)))))). Time Defined. (* constant 4839 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_n2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_real)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n))))))). Time Defined. (* constant 4840 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_natintrl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n)))))))). Time Defined. (* constant 4841 *) Definition l_e_st_eq_landau_n_rt_rp_r_shiftr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_real)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) l_e_st_eq_landau_n_rt_rp_r_1rl)))))). Time Defined. (* constant 4842 *) Definition l_e_st_eq_landau_n_rt_rp_r_intshiftr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_intmn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_t8 x ix y iy ly n) y iy) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1)))))). Time Defined. (* constant 4843 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t8a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) y) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) l_e_st_eq_landau_n_rt_rp_r_1rl) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) y (l_e_st_eq_landau_n_rt_rp_r_plmn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_mnpl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y))))))). Time Defined. (* constant 4844 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t9a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m) l_e_st_eq_landau_n_rt_rp_r_1rl) y))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)) => l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m) l_e_st_eq_landau_n_rt_rp_r_1rl) y (l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m) l_e_st_eq_landau_n_rt_rp_r_1rl i))))))))). Time Defined. (* constant 4845 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t10a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly m))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly m) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_t8a x ix y iy ly n)) (l_e_st_eq_landau_n_rt_rp_r_shift_t9a x ix y iy ly n m i) (l_e_st_eq_landau_n_rt_rp_r_shift_t8a x ix y iy ly m))))))))). Time Defined. (* constant 4846 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t11a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly m))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)) => l_e_st_eq_landau_n_rt_rp_r_isntirl (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly m) (l_e_st_eq_landau_n_rt_rp_r_shift_t10a x ix y iy ly n m i))))))))). Time Defined. (* constant 4847 *) Definition l_e_st_eq_landau_n_rt_rp_r_iseshiftr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) n m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)) => l_e_st_eq_landau_n_isinne (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) n m (l_e_st_eq_landau_n_rt_rp_r_shift_t11a x ix y iy ly n m i))))))))). Time Defined. (* constant 4848 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) => l_e_st_eq_landau_n_rt_rp_r_satz188d (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x l_e_st_eq_landau_n_rt_rp_r_1rl m))))))). Time Defined. (* constant 4849 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) => l_e_st_eq_landau_n_rt_rp_r_ismore1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_plmn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_shift_t9 x ix y iy ly n m)))))))). Time Defined. (* constant 4850 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) => l_e_st_eq_landau_n_rt_rp_r_satz188d (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 y) (l_e_st_eq_landau_n_rt_rp_r_shift_t10 x ix y iy ly n m)))))))). Time Defined. (* constant 4851 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) => l_e_st_eq_landau_n_rt_rp_r_ismore1 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) y) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mnpl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t11 x ix y iy ly n m)))))))). Time Defined. (* constant 4852 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n)))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) => l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 x ix y iy ly))) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n))) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shift_t12 x ix y iy ly n m)))))))). Time Defined. (* constant 4853 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n))) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) => l_e_st_eq_landau_n_rt_rp_r_lemmaiva5 (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n))) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_shift_t13 x ix y iy ly n m)))))))). Time Defined. (* constant 4854 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t15 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) => l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n))) (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_isntrl2 (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n)) (l_e_st_eq_landau_n_rt_rp_r_shift_t14 x ix y iy ly n m)))))))). Time Defined. (* constant 4855 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t16 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_not (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_satz10d (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_shift_t7 x ix y iy ly n)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x) => l_e_st_eq_landau_n_rt_rp_r_shift_t15 x ix y iy ly n t))))))). Time Defined. (* constant 4856 *) Definition l_e_st_eq_landau_n_rt_rp_r_shiftrls : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_satz167e (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) x (l_e_st_eq_landau_n_rt_rp_r_shift_t16 x ix y iy ly n))))))). Time Defined. (* constant 4857 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t17 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl y l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)) => l_e_st_eq_landau_n_rt_rp_r_satz188d y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl m))))))). Time Defined. (* constant 4858 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t18 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_1rl y) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)) => l_e_st_eq_landau_n_rt_rp_r_ismore12 (l_e_st_eq_landau_n_rt_rp_r_pl y l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_1rl y) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) (l_e_st_eq_landau_n_rt_rp_r_compl y l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_plmn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_shift_t17 x ix y iy ly n m)))))))). Time Defined. (* constant 4859 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t19 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)), l_e_st_eq_landau_n_rt_rp_r_more l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)) => l_e_st_eq_landau_n_rt_rp_r_satz188a l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y (l_e_st_eq_landau_n_rt_rp_r_shift_t18 x ix y iy ly n m)))))))). Time Defined. (* constant 4860 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t20 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)), l_e_st_eq_landau_n_more l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)) => l_e_st_eq_landau_n_rt_rp_r_lemmaiva5 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_t19 x ix y iy ly n m)))))))). Time Defined. (* constant 4861 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t21 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_not (l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)) (l_e_st_eq_landau_n_more l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n)) (l_e_st_eq_landau_n_satz10d l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_more y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)) => l_e_st_eq_landau_n_rt_rp_r_shift_t20 x ix y iy ly n t))))))). Time Defined. (* constant 4862 *) Definition l_e_st_eq_landau_n_rt_rp_r_lsshiftr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_satz167e y (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_t21 x ix y iy ly n))))))). Time Defined. (* constant 4863 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t22 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_intrl u))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_and3e1 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x) a))))))). Time Defined. (* constant 4864 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t23 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_lessis y u))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_and3e2 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x) a))))))). Time Defined. (* constant 4865 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t24 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_lessis u x))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_and3e3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x) a))))))). Time Defined. (* constant 4866 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t25 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_satz188f (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 y) (l_e_st_eq_landau_n_rt_rp_r_satz188f u x l_e_st_eq_landau_n_rt_rp_r_1rl l))))))))). Time Defined. (* constant 4867 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t26 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is u x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is u x) => l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y (l_e_st_eq_landau_n_rt_rp_r_ispl1 u x l_e_st_eq_landau_n_rt_rp_r_1rl i))))))))). Time Defined. (* constant 4868 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t27 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_less u x) (l_e_st_eq_landau_n_rt_rp_r_is u x) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y)) (l_e_st_eq_landau_n_rt_rp_r_shift_t24 x ix y iy ly u a) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_shift_t25 x ix y iy ly u a t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is u x) => l_e_st_eq_landau_n_rt_rp_r_shift_t26 x ix y iy ly u a t)))))))). Time Defined. (* constant 4869 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_ul : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_nat))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_rt_rp_r_shiftl u (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly u a) y iy (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly u a)))))))). Time Defined. (* constant 4870 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t28 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_rt_rp_r_islessis12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 u (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly u a) y iy (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly u a))) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shift_t27 x ix y iy ly u a)))))))). Time Defined. (* constant 4871 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t29 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_not (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_imp_th3 (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly))) (l_e_st_eq_landau_n_rt_rp_r_satz167d (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shift_t28 x ix y iy ly u a)) (fun (t:l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_lemmaiva6 (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) t)))))))). Time Defined. (* constant 4872 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t30 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_satz10e (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_shift_t29 x ix y iy ly u a)))))))). Time Defined. (* constant 4873 *) Definition l_e_st_eq_landau_n_rt_rp_r_shiftl1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shift_t30 x ix y iy ly u a)))))))). Time Defined. (* constant 4874 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t31 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shift_t30 x ix y iy ly u a)))))))). Time Defined. (* constant 4875 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t32 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 u (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly u a) y iy (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly u a))) (l_e_st_eq_landau_n_rt_rp_r_isnterl (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_shift_t31 x ix y iy ly u a))))))))). Time Defined. (* constant 4876 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t33 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) y) l_e_st_eq_landau_n_rt_rp_r_1rl) u))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) y) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) l_e_st_eq_landau_n_rt_rp_r_1rl) u (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) y) (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_plmn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y)) (l_e_st_eq_landau_n_rt_rp_r_mnpl u l_e_st_eq_landau_n_rt_rp_r_1rl)))))))). Time Defined. (* constant 4877 *) Definition l_e_st_eq_landau_n_rt_rp_r_shiftinv1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_is u (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real u (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) y) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_shift_t33 x ix y iy ly u a) (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) y) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a)) y) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a)) y (l_e_st_eq_landau_n_rt_rp_r_shift_t32 x ix y iy ly u a)))))))))). Time Defined. (* constant 4878 *) Definition l_e_st_eq_landau_n_rt_rp_r_shiftinv2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a)) u))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real u (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a)) (l_e_st_eq_landau_n_rt_rp_r_shiftinv1 x ix y iy ly u a)))))))). Time Defined. (* constant 4879 *) Definition l_e_st_eq_landau_n_rt_rp_r_seq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (alpha:Type), Type)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (alpha:Type) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (v:l_e_st_eq_landau_n_rt_rp_r_lessis y t), (forall (w:l_e_st_eq_landau_n_rt_rp_r_lessis t x), alpha)))))))))). Time Defined. (* constant 4880 *) Definition l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (alpha:Type), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly alpha), Prop))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (alpha:Type) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly alpha) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (it:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (lt:l_e_st_eq_landau_n_rt_rp_r_lessis y t), (forall (tl:l_e_st_eq_landau_n_rt_rp_r_lessis t x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x), (forall (v:l_e_st_eq_landau_n_rt_rp_r_is t u), l_e_is alpha (s t it lt tl) (s u iu lu ul))))))))))))))))). Time Defined. (* constant 4881 *) Definition l_e_st_eq_landau_n_rt_rp_r_shiftf : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (alpha:Type), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly alpha), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), alpha)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (alpha:Type) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly alpha) => (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => f (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly t) (l_e_st_eq_landau_n_rt_rp_r_intshiftr x ix y iy ly t) (l_e_st_eq_landau_n_rt_rp_r_lsshiftr x ix y iy ly t) (l_e_st_eq_landau_n_rt_rp_r_shiftrls x ix y iy ly t))))))))). Time Defined. (* constant 4882 *) Definition l_e_st_eq_landau_n_rt_rp_r_inseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), Prop)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (v:l_e_st_eq_landau_n_rt_rp_r_lessis y t), (forall (w:l_e_st_eq_landau_n_rt_rp_r_lessis t x), l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl (s t u v w)) (l_e_st_eq_landau_n_rt_rp_r_lessis y (s t u v w)) (l_e_st_eq_landau_n_rt_rp_r_lessis (s t u v w) x))))))))))). Time Defined. (* constant 4883 *) Definition l_e_st_eq_landau_n_rt_rp_r_injseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), Prop)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (it:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (lt:l_e_st_eq_landau_n_rt_rp_r_lessis y t), (forall (tl:l_e_st_eq_landau_n_rt_rp_r_lessis t x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x), (forall (v:l_e_st_eq_landau_n_rt_rp_r_is (s t it lt tl) (s u iu lu ul)), l_e_st_eq_landau_n_rt_rp_r_is t u))))))))))))))). Time Defined. (* constant 4884 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_prop1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl v) (l_e_st_eq_landau_n_rt_rp_r_lessis y v) (l_e_st_eq_landau_n_rt_rp_r_lessis v x)), Prop))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl v) (l_e_st_eq_landau_n_rt_rp_r_lessis y v) (l_e_st_eq_landau_n_rt_rp_r_lessis v x)) => l_e_st_eq_landau_n_rt_rp_r_is u (s v (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly v a) (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly v a) (l_e_st_eq_landau_n_rt_rp_r_shift_t24 x ix y iy ly v a))))))))))). Time Defined. (* constant 4885 *) Definition l_e_st_eq_landau_n_rt_rp_r_improp : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), Prop)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl v) (l_e_st_eq_landau_n_rt_rp_r_lessis y v) (l_e_st_eq_landau_n_rt_rp_r_lessis v x)) (forall (t:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl v) (l_e_st_eq_landau_n_rt_rp_r_lessis y v) (l_e_st_eq_landau_n_rt_rp_r_lessis v x)), l_e_st_eq_landau_n_rt_rp_r_shift_prop1 x ix y iy ly s u v t))))))))). Time Defined. (* constant 4886 *) Definition l_e_st_eq_landau_n_rt_rp_r_imseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), Prop))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_some (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s u t)))))))). Time Defined. (* constant 4887 *) Definition l_e_st_eq_landau_n_rt_rp_r_surjseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), Prop)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (v:l_e_st_eq_landau_n_rt_rp_r_lessis y t), (forall (w:l_e_st_eq_landau_n_rt_rp_r_lessis t x), l_e_st_eq_landau_n_rt_rp_r_imseq x ix y iy ly s t)))))))))). Time Defined. (* constant 4888 *) Definition l_e_st_eq_landau_n_rt_rp_r_perm : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), Prop)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) (l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s))))))). Time Defined. (* constant 4889 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_ns : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_real)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_intshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_lsshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftrls x ix y iy ly n))))))))). Time Defined. (* constant 4890 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t34 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n)) (l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n)) (l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) x))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => ins (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_intshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_lsshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftrls x ix y iy ly n))))))))). Time Defined. (* constant 4891 *) Definition l_e_st_eq_landau_n_rt_rp_r_shiftseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins t) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins t))))))))). Time Defined. (* constant 4892 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t35 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins n)) (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins m))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins n)) (l_e_st_eq_landau_n_rt_rp_r_shift_t30 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins n)) (l_e_st_eq_landau_n_rt_rp_r_shift_ul x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins m)) (l_e_st_eq_landau_n_rt_rp_r_shift_t30 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins m)) i))))))))))). Time Defined. (* constant 4893 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t36 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) l_e_st_eq_landau_n_rt_rp_r_1rl) y)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)) => l_e_st_eq_landau_n_rt_rp_r_isrlint (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins n)) y iy (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins n))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) l_e_st_eq_landau_n_rt_rp_r_1rl) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins m)) y iy (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins m))) (l_e_st_eq_landau_n_rt_rp_r_shift_t35 x ix y iy ly s ins js n m i)))))))))))). Time Defined. (* constant 4894 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t37 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)) => l_e_st_eq_landau_n_rt_rp_r_satz188b (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_satz188b (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins m) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 y) (l_e_st_eq_landau_n_rt_rp_r_shift_t36 x ix y iy ly s ins js n m i))))))))))))). Time Defined. (* constant 4895 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t38 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)) => js (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_intshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_lsshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftrls x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly m) (l_e_st_eq_landau_n_rt_rp_r_intshiftr x ix y iy ly m) (l_e_st_eq_landau_n_rt_rp_r_lsshiftr x ix y iy ly m) (l_e_st_eq_landau_n_rt_rp_r_shiftrls x ix y iy ly m) (l_e_st_eq_landau_n_rt_rp_r_shift_t37 x ix y iy ly s ins js n m i)))))))))))). Time Defined. (* constant 4896 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t39 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly m)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)) => l_e_st_eq_landau_n_rt_rp_r_satz188b (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly m) y (l_e_st_eq_landau_n_rt_rp_r_satz188b (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly n) y) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_shift_n2 x ix y iy ly m) y) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_shift_t38 x ix y iy ly s ins js n m i))))))))))))). Time Defined. (* constant 4897 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t40 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly m)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)) => l_e_st_eq_landau_n_rt_rp_r_isntirl (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shift_n1 x ix y iy ly m) (l_e_st_eq_landau_n_rt_rp_r_shift_t39 x ix y iy ly s ins js n m i)))))))))))). Time Defined. (* constant 4898 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t41 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) n m))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins n) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins m)) => l_e_st_eq_landau_n_isinne (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly) n m (l_e_st_eq_landau_n_rt_rp_r_shift_t40 x ix y iy ly s ins js n m i)))))))))))). Time Defined. (* constant 4899 *) Definition l_e_st_eq_landau_n_rt_rp_r_injshiftseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s), l_e_injective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (js:l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) => (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (v:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins t) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins u)) => l_e_st_eq_landau_n_rt_rp_r_shift_t41 x ix y iy ly s ins js t u v))))))))))). Time Defined. (* constant 4900 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t42 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_imseq x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => ss (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_intshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_lsshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftrls x ix y iy ly n))))))))))). Time Defined. (* constant 4901 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t43 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u), l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u) => l_ande1 (l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) (forall (t:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_shift_prop1 x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u t) p)))))))))))). Time Defined. (* constant 4902 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t44 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (s u (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t24 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u) => l_r_ande2 (l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) (fun (t:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_rt_rp_r_shift_prop1 x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u t) p)))))))))))). Time Defined. (* constant 4903 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_ul1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u) => l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p))))))))))))). Time Defined. (* constant 4904 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t45 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u), l_e_st_eq_landau_n_rt_rp_r_is (s u (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t24 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p))) (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u) => pri u (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t24 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_intshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_lsshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shiftrls x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shiftinv1 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)))))))))))))). Time Defined. (* constant 4905 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t46 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u) => l_e_st_eq_landau_n_rt_rp_r_shiftinv1 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)))))))))))))). Time Defined. (* constant 4906 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t47 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) (s u (l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t24 x ix y iy ly u (l_e_st_eq_landau_n_rt_rp_r_shift_t43 x ix y iy ly s ins pri ss n u p))) (l_e_st_eq_landau_n_rt_rp_r_shift_ns x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p))) (l_e_st_eq_landau_n_rt_rp_r_shift_t44 x ix y iy ly s ins pri ss n u p) (l_e_st_eq_landau_n_rt_rp_r_shift_t45 x ix y iy ly s ins pri ss n u p) (l_e_st_eq_landau_n_rt_rp_r_shift_t46 x ix y iy ly s ins pri ss n u p))))))))))))). Time Defined. (* constant 4907 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t48 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) n (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u) => l_e_st_eq_landau_n_rt_rp_r_iseshiftr x ix y iy ly n (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p)) (l_e_st_eq_landau_n_rt_rp_r_shift_t47 x ix y iy ly s ins pri ss n u p))))))))))))). Time Defined. (* constant 4908 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t49 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u), l_e_image (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins) n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) u) => l_somei (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) n (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins t)) (l_e_st_eq_landau_n_rt_rp_r_shift_ul1 x ix y iy ly s ins pri ss n u p) (l_e_st_eq_landau_n_rt_rp_r_shift_t48 x ix y iy ly s ins pri ss n u p))))))))))))). Time Defined. (* constant 4909 *) Definition l_e_st_eq_landau_n_rt_rp_r_shift_t50 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)), l_e_image (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins) n)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_someapp l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) t) (l_e_st_eq_landau_n_rt_rp_r_shift_t42 x ix y iy ly s ins pri ss n) (l_e_image (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins) n) (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_improp x ix y iy ly s (l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n) t) => l_e_st_eq_landau_n_rt_rp_r_shift_t49 x ix y iy ly s ins pri ss n t u)))))))))))). Time Defined. (* constant 4910 *) Definition l_e_st_eq_landau_n_rt_rp_r_surjshiftseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s), l_e_surjective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ss:l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) => (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_shift_t50 x ix y iy ly s ins pri ss t)))))))))). Time Defined. (* constant 4911 *) Definition l_e_st_eq_landau_n_rt_rp_r_bijshiftseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => l_andi (l_e_injective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins)) (l_e_surjective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly)) (l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins)) (l_e_st_eq_landau_n_rt_rp_r_injshiftseq x ix y iy ly s ins (l_ande1 (l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) (l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) ps)) (l_e_st_eq_landau_n_rt_rp_r_surjshiftseq x ix y iy ly s ins pri (l_ande2 (l_e_st_eq_landau_n_rt_rp_r_injseq x ix y iy ly s) (l_e_st_eq_landau_n_rt_rp_r_surjseq x ix y iy ly s) ps))))))))))). Time Defined. (* constant 4912 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_complex : Type. exact (l_e_st_eq_landau_n_pair1type l_e_st_eq_landau_n_rt_rp_r_real). Time Defined. (* constant 4913 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_cx : Type. exact l_e_st_eq_landau_n_rt_rp_r_c_complex. Time Defined. (* constant 4914 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_is : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), Prop)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_is l_e_st_eq_landau_n_rt_rp_r_c_cx x y)). Time Defined. (* constant 4915 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_nis : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), Prop)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_not (l_e_st_eq_landau_n_rt_rp_r_c_is x y))). Time Defined. (* constant 4916 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_some : (forall (p:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_complex), Prop)), Prop). exact (fun (p:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_complex), Prop)) => l_some l_e_st_eq_landau_n_rt_rp_r_c_cx p). Time Defined. (* constant 4917 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_all : (forall (p:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_complex), Prop)), Prop). exact (fun (p:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_complex), Prop)) => l_all l_e_st_eq_landau_n_rt_rp_r_c_cx p). Time Defined. (* constant 4918 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_one : (forall (p:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_complex), Prop)), Prop). exact (fun (p:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_complex), Prop)) => l_e_one l_e_st_eq_landau_n_rt_rp_r_c_cx p). Time Defined. (* constant 4919 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_pli : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_complex)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_pair1 l_e_st_eq_landau_n_rt_rp_r_real a b)). Time Defined. (* constant 4920 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_re : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_first1 l_e_st_eq_landau_n_rt_rp_r_real x). Time Defined. (* constant 4921 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_im : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_second1 l_e_st_eq_landau_n_rt_rp_r_real x). Time Defined. (* constant 4922 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_reis : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) a)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_first1is1 l_e_st_eq_landau_n_rt_rp_r_real a b)). Time Defined. (* constant 4923 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isre : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is a (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_first1is2 l_e_st_eq_landau_n_rt_rp_r_real a b)). Time Defined. (* constant 4924 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_imis : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) b)). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_second1is1 l_e_st_eq_landau_n_rt_rp_r_real a b)). Time Defined. (* constant 4925 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isim : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is b (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_second1is2 l_e_st_eq_landau_n_rt_rp_r_real a b)). Time Defined. (* constant 4926 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_pliis : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_pair1is1 l_e_st_eq_landau_n_rt_rp_r_real x). Time Defined. (* constant 4927 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ispli : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_pair1is2 l_e_st_eq_landau_n_rt_rp_r_real x). Time Defined. (* constant 4928 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscere : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_re t) x y i))). Time Defined. (* constant 4929 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isceim : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_im t) x y i))). Time Defined. (* constant 4930 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isrecx1 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is a b), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli a c) (l_e_st_eq_landau_n_rt_rp_r_c_pli b c))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is a b) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_pli t c) a b i)))). Time Defined. (* constant 4931 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is a b), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli c a) (l_e_st_eq_landau_n_rt_rp_r_c_pli c b))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is a b) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_pli c t) a b i)))). Time Defined. (* constant 4932 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is a b), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is c d), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli a c) (l_e_st_eq_landau_n_rt_rp_r_c_pli b d))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is a b) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is c d) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pli a c) (l_e_st_eq_landau_n_rt_rp_r_c_pli b c) (l_e_st_eq_landau_n_rt_rp_r_c_pli b d) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx1 a b c i) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 c d b j))))))). Time Defined. (* constant 4933 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz206 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x). Time Defined. (* constant 4934 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz207 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx x y i))). Time Defined. (* constant 4935 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz208 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is y z), l_e_st_eq_landau_n_rt_rp_r_c_is x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is y z) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx x y z i j))))). Time Defined. (* constant 4936 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_0c : l_e_st_eq_landau_n_rt_rp_r_c_complex. exact (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0). Time Defined. (* constant 4937 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_1c : l_e_st_eq_landau_n_rt_rp_r_c_complex. exact (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0). Time Defined. (* constant 4938 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_pl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_complex)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). Time Defined. (* constant 4939 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_plis12a : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl a c) (l_e_st_eq_landau_n_rt_rp_r_pl b d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli c d))) (l_e_st_eq_landau_n_rt_rp_r_pl a c) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli c d))) (l_e_st_eq_landau_n_rt_rp_r_pl b d) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) a (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) c (l_e_st_eq_landau_n_rt_rp_r_c_reis a b) (l_e_st_eq_landau_n_rt_rp_r_c_reis c d)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) b (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) d (l_e_st_eq_landau_n_rt_rp_r_c_imis a b) (l_e_st_eq_landau_n_rt_rp_r_c_imis c d)))))). Time Defined. (* constant 4940 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_plis12b : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl a c) (l_e_st_eq_landau_n_rt_rp_r_pl b d)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl a c) (l_e_st_eq_landau_n_rt_rp_r_pl b d)) (l_e_st_eq_landau_n_rt_rp_r_c_plis12a a b c d))))). Time Defined. (* constant 4941 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_plis1a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_pl s (l_e_st_eq_landau_n_rt_rp_r_c_im x)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_pl s (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) r (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_reis r s)) (l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) s (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_imis r s))))). Time Defined. (* constant 4942 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_plis1b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_pl s (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_pl s (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_plis1a x r s)))). Time Defined. (* constant 4943 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_plis2a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) s))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) s) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) r (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_reis r s)) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) s (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_imis r s))))). Time Defined. (* constant 4944 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_plis2b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) s)) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) s)) (l_e_st_eq_landau_n_rt_rp_r_c_plis2a x r s)))). Time Defined. (* constant 4945 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ispl1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl t z) x y i)))). Time Defined. (* constant 4946 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ispl2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl z x) (l_e_st_eq_landau_n_rt_rp_r_c_pl z y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl z t) x y i)))). Time Defined. (* constant 4947 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ispl12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z u), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y u))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z u) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) (l_e_st_eq_landau_n_rt_rp_r_c_ispl1 x y z i) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 z u y j))))))). Time Defined. (* constant 4948 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz209 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). Time Defined. (* constant 4949 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_compl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz209 x y)). Time Defined. (* constant 4950 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz210 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x l_e_st_eq_landau_n_rt_rp_r_c_0c) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) x (l_e_st_eq_landau_n_rt_rp_r_c_plis2a x l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_pliis x)). Time Defined. (* constant 4951 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz210a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_pl x l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl x l_e_st_eq_landau_n_rt_rp_r_c_0c) x (l_e_st_eq_landau_n_rt_rp_r_c_satz210 x)). Time Defined. (* constant 4952 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz210b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl l_e_st_eq_landau_n_rt_rp_r_c_0c x) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl l_e_st_eq_landau_n_rt_rp_r_c_0c x) (l_e_st_eq_landau_n_rt_rp_r_c_pl x l_e_st_eq_landau_n_rt_rp_r_c_0c) x (l_e_st_eq_landau_n_rt_rp_r_c_compl l_e_st_eq_landau_n_rt_rp_r_c_0c x) (l_e_st_eq_landau_n_rt_rp_r_c_satz210 x)). Time Defined. (* constant 4953 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz210c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_pl l_e_st_eq_landau_n_rt_rp_r_c_0c x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl l_e_st_eq_landau_n_rt_rp_r_c_0c x) x (l_e_st_eq_landau_n_rt_rp_r_c_satz210b x)). Time Defined. (* constant 4954 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz211 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_plis1a z (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_asspl1 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_asspl1 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_c_plis2b x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))))). Time Defined. (* constant 4955 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_asspl1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz211 x y z))). Time Defined. (* constant 4956 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_asspl2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_asspl1 x y z)))). Time Defined. (* constant 4957 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2212_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re u)) (l_e_st_eq_landau_n_rt_rp_r_c_re x))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re u)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pl y u)) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_isre (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re u)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im u))) (l_e_st_eq_landau_n_rt_rp_r_c_iscere (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x i))))). Time Defined. (* constant 4958 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2212_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im u)) (l_e_st_eq_landau_n_rt_rp_r_c_im x))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im u)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pl y u)) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_isim (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re u)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im u))) (l_e_st_eq_landau_n_rt_rp_r_c_isceim (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x i))))). Time Defined. (* constant 4959 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2212_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re u) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x) => l_e_st_eq_landau_n_rt_rp_r_satz187d (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re u) (l_e_st_eq_landau_n_rt_rp_r_c_2212_t1 x y u i))))). Time Defined. (* constant 4960 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2212_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im u) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x) => l_e_st_eq_landau_n_rt_rp_r_satz187d (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im u) (l_e_st_eq_landau_n_rt_rp_r_c_2212_t2 x y u i))))). Time Defined. (* constant 4961 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x), l_e_st_eq_landau_n_rt_rp_r_c_is u (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx u (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re u) (l_e_st_eq_landau_n_rt_rp_r_c_im u)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_ispli u) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_re u) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im u) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_2212_t3 x y u i) (l_e_st_eq_landau_n_rt_rp_r_c_2212_t4 x y u i)))))). Time Defined. (* constant 4962 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) x (l_e_st_eq_landau_n_rt_rp_r_c_plis2a y (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_satz187a (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_satz187a (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_pliis x))). Time Defined. (* constant 4963 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_some (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y t) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_somei l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y t) x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_satz212b x y))). Time Defined. (* constant 4964 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_one (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y t) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_onei l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y t) x) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y t) x) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx t u (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_satz212a x y t v) (l_e_st_eq_landau_n_rt_rp_r_c_satz212a x y u w))))) (l_e_st_eq_landau_n_rt_rp_r_c_satz212c x y))). Time Defined. (* constant 4965 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mn : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_complex)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). Time Defined. (* constant 4966 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mnis12a : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn a c) (l_e_st_eq_landau_n_rt_rp_r_mn b d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli c d))) (l_e_st_eq_landau_n_rt_rp_r_mn a c) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli c d))) (l_e_st_eq_landau_n_rt_rp_r_mn b d) (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) a (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) c (l_e_st_eq_landau_n_rt_rp_r_c_reis a b) (l_e_st_eq_landau_n_rt_rp_r_c_reis c d)) (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) b (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) d (l_e_st_eq_landau_n_rt_rp_r_c_imis a b) (l_e_st_eq_landau_n_rt_rp_r_c_imis c d)))))). Time Defined. (* constant 4967 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mnis12b : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn a c) (l_e_st_eq_landau_n_rt_rp_r_mn b d)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn a c) (l_e_st_eq_landau_n_rt_rp_r_mn b d)) (l_e_st_eq_landau_n_rt_rp_r_c_mnis12a a b c d))))). Time Defined. (* constant 4968 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mnis1a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_mn s (l_e_st_eq_landau_n_rt_rp_r_c_im x)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_mn r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_mn s (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) r (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_reis r s)) (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) s (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_imis r s))))). Time Defined. (* constant 4969 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mnis1b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_mn s (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_mn s (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_mnis1a x r s)))). Time Defined. (* constant 4970 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mnis2a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) s))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) s) (l_e_st_eq_landau_n_rt_rp_r_ismn2 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) r (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_reis r s)) (l_e_st_eq_landau_n_rt_rp_r_ismn2 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) s (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_imis r s))))). Time Defined. (* constant 4971 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mnis2b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) s)) (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) s)) (l_e_st_eq_landau_n_rt_rp_r_c_mnis2a x r s)))). Time Defined. (* constant 4972 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ismn1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) (l_e_st_eq_landau_n_rt_rp_r_c_mn y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_mn t z) x y i)))). Time Defined. (* constant 4973 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ismn2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn z x) (l_e_st_eq_landau_n_rt_rp_r_c_mn z y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_mn z t) x y i)))). Time Defined. (* constant 4974 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ismn12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z u), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) (l_e_st_eq_landau_n_rt_rp_r_c_mn y u))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z u) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) (l_e_st_eq_landau_n_rt_rp_r_c_mn y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn y u) (l_e_st_eq_landau_n_rt_rp_r_c_ismn1 x y z i) (l_e_st_eq_landau_n_rt_rp_r_c_ismn2 z u y j))))))). Time Defined. (* constant 4975 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212d : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x), l_e_st_eq_landau_n_rt_rp_r_c_is u (l_e_st_eq_landau_n_rt_rp_r_c_mn x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x) => l_e_st_eq_landau_n_rt_rp_r_c_satz212a x y u i)))). Time Defined. (* constant 4976 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212e : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) u)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) x) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx u (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz212d x y u i))))). Time Defined. (* constant 4977 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212f : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl u y) x), l_e_st_eq_landau_n_rt_rp_r_c_is u (l_e_st_eq_landau_n_rt_rp_r_c_mn x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl u y) x) => l_e_st_eq_landau_n_rt_rp_r_c_satz212d x y u (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) (l_e_st_eq_landau_n_rt_rp_r_c_pl u y) x (l_e_st_eq_landau_n_rt_rp_r_c_compl y u) i))))). Time Defined. (* constant 4978 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212g : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl u y) x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) u)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl u y) x) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx u (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz212f x y u i))))). Time Defined. (* constant 4979 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz212h : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212b x y)). Time Defined. (* constant 4980 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2213_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_re l_e_st_eq_landau_n_rt_rp_r_c_0c) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_isre (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_iscere (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c i) (l_e_st_eq_landau_n_rt_rp_r_c_reis l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 4981 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2213_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_im l_e_st_eq_landau_n_rt_rp_r_c_0c) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_isim (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_isceim (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c i) (l_e_st_eq_landau_n_rt_rp_r_c_imis l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 4982 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2213_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_satz182b (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_2213_t1 x y i)))). Time Defined. (* constant 4983 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2213_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_satz182b (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_2213_t2 x y i)))). Time Defined. (* constant 4984 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz213a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) y (l_e_st_eq_landau_n_rt_rp_r_c_ispli x) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_2213_t3 x y i) (l_e_st_eq_landau_n_rt_rp_r_c_2213_t4 x y i)) (l_e_st_eq_landau_n_rt_rp_r_c_pliis y)))). Time Defined. (* constant 4985 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2213_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_st_eq_landau_n_rt_rp_r_satz182e (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_iscere x y i)))). Time Defined. (* constant 4986 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2213_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_st_eq_landau_n_rt_rp_r_satz182e (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_isceim x y i)))). Time Defined. (* constant 4987 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz213b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) l_e_st_eq_landau_n_rt_rp_r_c_0c))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_2213_t5 x y i) (l_e_st_eq_landau_n_rt_rp_r_c_2213_t6 x y i)))). Time Defined. (* constant 4988 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_m0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_complex). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_mn l_e_st_eq_landau_n_rt_rp_r_c_0c x). Time Defined. (* constant 4989 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz214 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_im l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_pl01 (l_e_st_eq_landau_n_rt_rp_r_c_re l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_reis l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl01 (l_e_st_eq_landau_n_rt_rp_r_c_im l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_imis l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 4990 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz214a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_satz214 x)). Time Defined. (* constant 4991 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_m0isa : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 a) (l_e_st_eq_landau_n_rt_rp_r_m0 b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 a) (l_e_st_eq_landau_n_rt_rp_r_m0 b)) (l_e_st_eq_landau_n_rt_rp_r_c_satz214 (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b))) (l_e_st_eq_landau_n_rt_rp_r_m0 a) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b))) (l_e_st_eq_landau_n_rt_rp_r_m0 b) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) a (l_e_st_eq_landau_n_rt_rp_r_c_reis a b)) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) b (l_e_st_eq_landau_n_rt_rp_r_c_imis a b))))). Time Defined. (* constant 4992 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_m0isb : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 a) (l_e_st_eq_landau_n_rt_rp_r_m0 b)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 a) (l_e_st_eq_landau_n_rt_rp_r_m0 b)) (l_e_st_eq_landau_n_rt_rp_r_c_m0isa a b))). Time Defined. (* constant 4993 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ism0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_m0 t) x y i))). Time Defined. (* constant 4994 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz215 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) x (l_e_st_eq_landau_n_rt_rp_r_c_ism0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_satz214 x)) (l_e_st_eq_landau_n_rt_rp_r_c_m0isa (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_satz177 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_satz177 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_pliis x)). Time Defined. (* constant 4995 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz215a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) x (l_e_st_eq_landau_n_rt_rp_r_c_satz215 x)). Time Defined. (* constant 4996 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz215b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) y (l_e_st_eq_landau_n_rt_rp_r_c_ism0 x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y) i) (l_e_st_eq_landau_n_rt_rp_r_c_satz215 y)))). Time Defined. (* constant 4997 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz215c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)), l_e_st_eq_landau_n_rt_rp_r_c_is y (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y (l_e_st_eq_landau_n_rt_rp_r_c_satz215b x y i)))). Time Defined. (* constant 4998 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz215d : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) => l_e_st_eq_landau_n_rt_rp_r_c_satz215c y x (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y i)))). Time Defined. (* constant 4999 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz215e : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 y) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y) (l_e_st_eq_landau_n_rt_rp_r_c_satz215d x y i)))). Time Defined. (* constant 5000 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz216 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) l_e_st_eq_landau_n_rt_rp_r_c_0c). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) x (l_e_st_eq_landau_n_rt_rp_r_c_satz214 x)) (l_e_st_eq_landau_n_rt_rp_r_c_plis2a x (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satz179 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_satz179 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). Time Defined. (* constant 5001 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2216_anders : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) l_e_st_eq_landau_n_rt_rp_r_c_0c). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212h l_e_st_eq_landau_n_rt_rp_r_c_0c x). Time Defined. (* constant 5002 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz216a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) x) l_e_st_eq_landau_n_rt_rp_r_c_0c). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) x) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_compl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) x) (l_e_st_eq_landau_n_rt_rp_r_c_satz216 x)). Time Defined. (* constant 5003 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz217 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_m0isa (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_satz180 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_satz180 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_plis12b (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y) (l_e_st_eq_landau_n_rt_rp_r_c_satz214a x) (l_e_st_eq_landau_n_rt_rp_r_c_satz214a y)))). Time Defined. (* constant 5004 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz217a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz217 x y))). Time Defined. (* constant 5005 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz218 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_plis2b x (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y) x (l_e_st_eq_landau_n_rt_rp_r_c_satz214a y)))). Time Defined. (* constant 5006 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz218a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz218 x y))). Time Defined. (* constant 5007 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_2219_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 y))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_ism0 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz218 x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz217 x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) y (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_satz215 y)))). Time Defined. (* constant 5008 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz219 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_mn y x) (l_e_st_eq_landau_n_rt_rp_r_c_2219_t1 x y) (l_e_st_eq_landau_n_rt_rp_r_c_compl (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_satz218a y x))). Time Defined. (* constant 5009 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz219a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mn y x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mn y x)) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz219 y x))). Time Defined. (* constant 5010 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ts : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_complex)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))))). Time Defined. (* constant 5011 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_rets : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). Time Defined. (* constant 5012 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_imts : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))). Time Defined. (* constant 5013 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rets (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts a c) (l_e_st_eq_landau_n_rt_rp_r_ts b d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli c d))) (l_e_st_eq_landau_n_rt_rp_r_ts a c) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli c d))) (l_e_st_eq_landau_n_rt_rp_r_ts b d) (l_e_st_eq_landau_n_rt_rp_r_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) a (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) c (l_e_st_eq_landau_n_rt_rp_r_c_reis a b) (l_e_st_eq_landau_n_rt_rp_r_c_reis c d)) (l_e_st_eq_landau_n_rt_rp_r_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) b (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) d (l_e_st_eq_landau_n_rt_rp_r_c_imis a b) (l_e_st_eq_landau_n_rt_rp_r_c_imis c d)))))). Time Defined. (* constant 5014 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_imts (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts a d) (l_e_st_eq_landau_n_rt_rp_r_ts b c)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli c d))) (l_e_st_eq_landau_n_rt_rp_r_ts a d) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli c d))) (l_e_st_eq_landau_n_rt_rp_r_ts b c) (l_e_st_eq_landau_n_rt_rp_r_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) a (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) d (l_e_st_eq_landau_n_rt_rp_r_c_reis a b) (l_e_st_eq_landau_n_rt_rp_r_c_imis c d)) (l_e_st_eq_landau_n_rt_rp_r_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) b (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) c (l_e_st_eq_landau_n_rt_rp_r_c_imis a b) (l_e_st_eq_landau_n_rt_rp_r_c_reis c d)))))). Time Defined. (* constant 5015 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_tsis12a : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts a c) (l_e_st_eq_landau_n_rt_rp_r_ts b d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts a d) (l_e_st_eq_landau_n_rt_rp_r_ts b c))))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_rets (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts a c) (l_e_st_eq_landau_n_rt_rp_r_ts b d)) (l_e_st_eq_landau_n_rt_rp_r_c_imts (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts a d) (l_e_st_eq_landau_n_rt_rp_r_ts b c)) (l_e_st_eq_landau_n_rt_rp_r_c_v3_t1 a b c d) (l_e_st_eq_landau_n_rt_rp_r_c_v3_t2 a b c d))))). Time Defined. (* constant 5016 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_tsis12b : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts a c) (l_e_st_eq_landau_n_rt_rp_r_ts b d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts a d) (l_e_st_eq_landau_n_rt_rp_r_ts b c))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli a b) (l_e_st_eq_landau_n_rt_rp_r_c_pli c d)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts a c) (l_e_st_eq_landau_n_rt_rp_r_ts b d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts a d) (l_e_st_eq_landau_n_rt_rp_r_ts b c))) (l_e_st_eq_landau_n_rt_rp_r_c_tsis12a a b c d))))). Time Defined. (* constant 5017 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rets (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_im x)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) r (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_reis r s)) (l_e_st_eq_landau_n_rt_rp_r_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) s (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_imis r s))))). Time Defined. (* constant 5018 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_imts (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_re x)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) r (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_reis r s)) (l_e_st_eq_landau_n_rt_rp_r_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) s (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_imis r s))))). Time Defined. (* constant 5019 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_tsis1a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_re x))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_rets (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_imts (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_c_v3_t3 x r s) (l_e_st_eq_landau_n_rt_rp_r_c_v3_t4 x r s)))). Time Defined. (* constant 5020 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_tsis1b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_re x)))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_c_re x)))) (l_e_st_eq_landau_n_rt_rp_r_c_tsis1a x r s)))). Time Defined. (* constant 5021 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rets x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) s))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) s) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) r (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_reis r s)) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) s (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_imis r s))))). Time Defined. (* constant 5022 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_imts x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) r))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) r) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) s (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_imis r s)) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) r (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_reis r s))))). Time Defined. (* constant 5023 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_tsis2a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) r)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_rets x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) s)) (l_e_st_eq_landau_n_rt_rp_r_c_imts x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) r)) (l_e_st_eq_landau_n_rt_rp_r_c_v3_t5 x r s) (l_e_st_eq_landau_n_rt_rp_r_c_v3_t6 x r s)))). Time Defined. (* constant 5024 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_tsis2b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) r))) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) s)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) s) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) r))) (l_e_st_eq_landau_n_rt_rp_r_c_tsis2a x r s)))). Time Defined. (* constant 5025 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ists1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t z) x y i)))). Time Defined. (* constant 5026 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ists2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts z t) x y i)))). Time Defined. (* constant 5027 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ists12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z u), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z u) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 x y z i) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 z u y j))))))). Time Defined. (* constant 5028 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3220_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_rets y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). Time Defined. (* constant 5029 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3220_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_imts y x) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))))). Time Defined. (* constant 5030 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz220 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_rets y x) (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts y x) (l_e_st_eq_landau_n_rt_rp_r_c_3220_t1 x y) (l_e_st_eq_landau_n_rt_rp_r_c_3220_t2 x y))). Time Defined. (* constant 5031 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_comts : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz220 x y)). Time Defined. (* constant 5032 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re l_e_st_eq_landau_n_rt_rp_r_c_0c) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_iscere x l_e_st_eq_landau_n_rt_rp_r_c_0c i) (l_e_st_eq_landau_n_rt_rp_r_c_reis l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 5033 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im l_e_st_eq_landau_n_rt_rp_r_c_0c) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_isceim x l_e_st_eq_landau_n_rt_rp_r_c_0c i) (l_e_st_eq_landau_n_rt_rp_r_c_imis l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 5034 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mod2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))). Time Defined. (* constant 5035 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma1 x i)) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma2 x i))) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 5036 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_re2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)). Time Defined. (* constant 5037 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_im2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)). Time Defined. (* constant 5038 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) n (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_ispli x) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0 i t))))). Time Defined. (* constant 5039 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ispos (l_e_st_eq_landau_n_rt_rp_r_c_v3_im2 x) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_v3_im2 x) (l_e_st_eq_landau_n_rt_rp_r_pl01 (l_e_st_eq_landau_n_rt_rp_r_c_v3_re2 x) (l_e_st_eq_landau_n_rt_rp_r_c_v3_im2 x) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x) i))) (l_e_st_eq_landau_n_rt_rp_r_possq (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_v3_t7 x n i))))). Time Defined. (* constant 5040 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ispos (l_e_st_eq_landau_n_rt_rp_r_c_v3_re2 x) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_v3_re2 x) (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_c_v3_re2 x) (l_e_st_eq_landau_n_rt_rp_r_c_v3_im2 x) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x) i))) (l_e_st_eq_landau_n_rt_rp_r_possq (l_e_st_eq_landau_n_rt_rp_r_c_re x) o))))). Time Defined. (* constant 5041 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0), (forall (p:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_pospl (l_e_st_eq_landau_n_rt_rp_r_c_v3_re2 x) (l_e_st_eq_landau_n_rt_rp_r_c_v3_im2 x) (l_e_st_eq_landau_n_rt_rp_r_possq (l_e_st_eq_landau_n_rt_rp_r_c_re x) o) (l_e_st_eq_landau_n_rt_rp_r_possq (l_e_st_eq_landau_n_rt_rp_r_c_im x) p))))). Time Defined. (* constant 5042 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v3_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_v3_t9 x n o t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_v3_t10 x n o t)))). Time Defined. (* constant 5043 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_v3_t8 x n t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_v3_t11 x n t))). Time Defined. (* constant 5044 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_0notn (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma3 x t)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_pnotn (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma4 x t))). Time Defined. (* constant 5045 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma1 x i))). Time Defined. (* constant 5046 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma2 x i))). Time Defined. (* constant 5047 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t1 x y i)) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t2 x y i))) (l_e_st_eq_landau_n_rt_rp_r_satz187c l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl02 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))))). Time Defined. (* constant 5048 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t1 x y i)) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t2 x y i))) (l_e_st_eq_landau_n_rt_rp_r_pl02 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 5049 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz221a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_3221_t3 x y i) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t4 x y i)))). Time Defined. (* constant 5050 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz221b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y x) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_comts x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz221a y x i)))). Time Defined. (* constant 5051 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma4 y n)))). Time Defined. (* constant 5052 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_reis (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma1 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) i))))). Time Defined. (* constant 5053 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_imis (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma2 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) i))))). Time Defined. (* constant 5054 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t6 x y i n)) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t7 x y i n))) (l_e_st_eq_landau_n_rt_rp_r_pl02 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))))). Time Defined. (* constant 5055 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_ii1r : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))). Time Defined. (* constant 5056 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ir : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))). Time Defined. (* constant 5057 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_ir1i : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))). Time Defined. (* constant 5058 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). Time Defined. (* constant 5059 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_rr1r : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))). Time Defined. (* constant 5060 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))). Time Defined. (* constant 5061 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_ri1r : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))). Time Defined. (* constant 5062 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ir : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))). Time Defined. (* constant 5063 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_ri1i : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))). Time Defined. (* constant 5064 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))). Time Defined. (* constant 5065 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_3221_ii1r x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_3221_ii1r x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ir x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y) (l_e_st_eq_landau_n_rt_rp_r_assts1 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))))))). Time Defined. (* constant 5066 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3221_rr1r x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_ii1r x y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_disttm1 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_3221_rr1r x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_ii1r x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y) (l_e_st_eq_landau_n_rt_rp_r_assts1 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t9 x y i n)))))). Time Defined. (* constant 5067 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3221_ri1i x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_ir1i x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y)) (l_e_st_eq_landau_n_rt_rp_r_disttp1 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_3221_ri1i x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_ir1i x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y) (l_e_st_eq_landau_n_rt_rp_r_assts1 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_assts1 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)))))). Time Defined. (* constant 5068 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t10 x y i n) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t11 x y i n)) (l_e_st_eq_landau_n_rt_rp_r_asspl2 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y)) (l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1ii x y) (l_e_st_eq_landau_n_rt_rp_r_plmn (l_e_st_eq_landau_n_rt_rp_r_c_3221_r1rr x y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_i1ri x y))) (l_e_st_eq_landau_n_rt_rp_r_distpt2 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))))))). Time Defined. (* constant 5069 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t12 x y i n) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t8 x y i n))))). Time Defined. (* constant 5070 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz192c (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t13 x y i n)) (l_e_st_eq_landau_n_rt_rp_r_pnot0 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t5 x y i n)))))). Time Defined. (* constant 5071 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t15 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_satz182b (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t6 x y i n)) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t14 x y i n)))))). Time Defined. (* constant 5072 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t16 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_pl01 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t14 x y i n))) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t7 x y i n))))). Time Defined. (* constant 5073 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t17 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_im y) l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_v3_t7 y n j))))). Time Defined. (* constant 5074 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t18 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im y) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz192c (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t15 x y i n)) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t17 x y i n j)))))). Time Defined. (* constant 5075 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t19 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz192c (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t16 x y i n)) o))))). Time Defined. (* constant 5076 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t20 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_3221_t18 x y i n t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_re y) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_3221_t19 x y i n t))))). Time Defined. (* constant 5077 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3221_t21 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_ispli x) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_3221_t14 x y i n) (l_e_st_eq_landau_n_rt_rp_r_c_3221_t20 x y i n)))))). Time Defined. (* constant 5078 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz221c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is y l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_or_th2 (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is y l_e_st_eq_landau_n_rt_rp_r_c_0c) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_3221_t21 x y i t)))). Time Defined. (* constant 5079 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz221d : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_or (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_or_th3 (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is y l_e_st_eq_landau_n_rt_rp_r_c_0c) n o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz221c x y t))))). Time Defined. (* constant 5080 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3222_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_re x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satz195 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts02 (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz176b l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 5081 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3222_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_c_im x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_ts02 (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_satz195 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 5082 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz222 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x l_e_st_eq_landau_n_rt_rp_r_c_1c) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_1rl))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) x (l_e_st_eq_landau_n_rt_rp_r_c_tsis2a x l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_3222_t1 x) (l_e_st_eq_landau_n_rt_rp_r_c_3222_t2 x)) (l_e_st_eq_landau_n_rt_rp_r_c_pliis x)). Time Defined. (* constant 5083 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz222a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_ts x l_e_st_eq_landau_n_rt_rp_r_c_1c)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x l_e_st_eq_landau_n_rt_rp_r_c_1c) x (l_e_st_eq_landau_n_rt_rp_r_c_satz222 x)). Time Defined. (* constant 5084 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz222b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c x) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c x) (l_e_st_eq_landau_n_rt_rp_r_c_ts x l_e_st_eq_landau_n_rt_rp_r_c_1c) x (l_e_st_eq_landau_n_rt_rp_r_c_comts l_e_st_eq_landau_n_rt_rp_r_c_1c x) (l_e_st_eq_landau_n_rt_rp_r_c_satz222 x)). Time Defined. (* constant 5085 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz222c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c x) x (l_e_st_eq_landau_n_rt_rp_r_c_satz222b x)). Time Defined. (* constant 5086 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3223_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_satz197b (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_satz195 (l_e_st_eq_landau_n_rt_rp_r_c_re x)))) (l_e_st_eq_landau_n_rt_rp_r_ts02 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz176b l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))) (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz176b l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 5087 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3223_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl))) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts02 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz176b l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_satz197b (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_satz195 (l_e_st_eq_landau_n_rt_rp_r_c_im x))))) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 5088 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz223 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) x (l_e_st_eq_landau_n_rt_rp_r_c_m0isa l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_tsis2a x (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_3223_t1 x) (l_e_st_eq_landau_n_rt_rp_r_c_3223_t2 x)) (l_e_st_eq_landau_n_rt_rp_r_c_satz214a x)). Time Defined. (* constant 5089 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz223a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_satz223 x)). Time Defined. (* constant 5090 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz223b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c) x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_comts (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c) x) (l_e_st_eq_landau_n_rt_rp_r_c_satz223 x)). Time Defined. (* constant 5091 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz223c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c) x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_satz223b x)). Time Defined. (* constant 5092 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))). Time Defined. (* constant 5093 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))). Time Defined. (* constant 5094 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3224_rxiy : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))). Time Defined. (* constant 5095 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3224_ixry : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))). Time Defined. (* constant 5096 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3224_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y) (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y))) (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y)) (l_e_st_eq_landau_n_rt_rp_r_satz197a (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_satz197a (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y))) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y)) (l_e_st_eq_landau_n_rt_rp_r_satz177 (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y))) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y)) (l_e_st_eq_landau_n_rt_rp_r_satz181a (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixiy x y) (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxry x y)))). Time Defined. (* constant 5097 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3224_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxiy x y) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixry x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxiy x y)) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixry x y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxiy x y) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixry x y))) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxiy x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixry x y)) (l_e_st_eq_landau_n_rt_rp_r_satz197a (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_satz197a (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_satz180a (l_e_st_eq_landau_n_rt_rp_r_c_3224_rxiy x y) (l_e_st_eq_landau_n_rt_rp_r_c_3224_ixry x y)))). Time Defined. (* constant 5098 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz224a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) y) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))))) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) y (l_e_st_eq_landau_n_rt_rp_r_c_satz214 x)) (l_e_st_eq_landau_n_rt_rp_r_c_tsis1a y (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_c_3224_t1 x y) (l_e_st_eq_landau_n_rt_rp_r_c_3224_t2 x y)) (l_e_st_eq_landau_n_rt_rp_r_c_m0isb (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))))). Time Defined. (* constant 5099 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz224b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 y) x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts y x)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_comts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz224a y x) (l_e_st_eq_landau_n_rt_rp_r_c_ism0 (l_e_st_eq_landau_n_rt_rp_r_c_ts y x) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_comts y x)))). Time Defined. (* constant 5100 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz224c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz224a x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz224b x y))). Time Defined. (* constant 5101 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz224d : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz224c x y))). Time Defined. (* constant 5102 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz224e : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) y) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz224a x y))). Time Defined. (* constant 5103 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz224f : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz224b x y))). Time Defined. (* constant 5104 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz225 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 y))) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz224c x (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) y x (l_e_st_eq_landau_n_rt_rp_r_c_satz215 y)))). Time Defined. (* constant 5105 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz225a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz225 x y))). Time Defined. (* constant 5106 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))). Time Defined. (* constant 5107 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_iir : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))). Time Defined. (* constant 5108 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_rii : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))). Time Defined. (* constant 5109 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_iri : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))). Time Defined. (* constant 5110 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_rri : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))). Time Defined. (* constant 5111 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_iii : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))). Time Defined. (* constant 5112 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_rir : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))). Time Defined. (* constant 5113 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_irr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))). Time Defined. (* constant 5114 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z))))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)) (l_e_st_eq_landau_n_rt_rp_r_disttm1 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_disttp1 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_asspl1 (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_satz180a (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z))))))). Time Defined. (* constant 5115 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_disttm1 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_disttp1 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z))) (l_e_st_eq_landau_n_rt_rp_r_asspl2 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)))))). Time Defined. (* constant 5116 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z))) (l_e_st_eq_landau_n_rt_rp_r_c_tsis1a z (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t2 x y z))))). Time Defined. (* constant 5117 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts a b) c) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts b c) a)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts a b) c) (l_e_st_eq_landau_n_rt_rp_r_ts a (l_e_st_eq_landau_n_rt_rp_r_ts b c)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_ts b c) a) (l_e_st_eq_landau_n_rt_rp_r_assts1 a b c) (l_e_st_eq_landau_n_rt_rp_r_comts a (l_e_st_eq_landau_n_rt_rp_r_ts b c))))). Time Defined. (* constant 5118 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t5 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) c) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl c a) b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) c) (l_e_st_eq_landau_n_rt_rp_r_pl c (l_e_st_eq_landau_n_rt_rp_r_pl a b)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl c a) b) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_pl a b) c) (l_e_st_eq_landau_n_rt_rp_r_asspl2 c a b)))). Time Defined. (* constant 5119 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t6 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl a (l_e_st_eq_landau_n_rt_rp_r_pl b c)) (l_e_st_eq_landau_n_rt_rp_r_pl b (l_e_st_eq_landau_n_rt_rp_r_pl c a))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl a (l_e_st_eq_landau_n_rt_rp_r_pl b c)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl b c) a) (l_e_st_eq_landau_n_rt_rp_r_pl b (l_e_st_eq_landau_n_rt_rp_r_pl c a)) (l_e_st_eq_landau_n_rt_rp_r_compl a (l_e_st_eq_landau_n_rt_rp_r_pl b c)) (l_e_st_eq_landau_n_rt_rp_r_asspl1 b c a)))). Time Defined. (* constant 5120 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr y z x))). Time Defined. (* constant 5121 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_iir y z x))). Time Defined. (* constant 5122 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_rii y z x))). Time Defined. (* constant 5123 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_iri y z x))). Time Defined. (* constant 5124 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_rri y z x))). Time Defined. (* constant 5125 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_iii y z x))). Time Defined. (* constant 5126 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_rir y z x))). Time Defined. (* constant 5127 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_irr y z x))). Time Defined. (* constant 5128 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z))) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t6 (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))))))). Time Defined. (* constant 5129 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t5 (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))))). Time Defined. (* constant 5130 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z))) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t7 x y z)))). Time Defined. (* constant 5131 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 x y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t8 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t4 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))))). Time Defined. (* constant 5132 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 x y z)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 x y z))) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t3 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii x y z)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t9 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t10 x y z))))). Time Defined. (* constant 5133 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 x y z)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 x y z))) (l_e_st_eq_landau_n_rt_rp_r_c_comts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t3 y z x)))). Time Defined. (* constant 5134 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3226_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_3226_rrr1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_iir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rii1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iri1 x y z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3226_rir1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_irr1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_rri1 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_3226_iii1 x y z))) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t11 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3226_t12 x y z)))). Time Defined. (* constant 5135 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz226 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3226_t13 x y z))). Time Defined. (* constant 5136 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_assts1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz226 x y z))). Time Defined. (* constant 5137 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_assts2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 x y z)))). Time Defined. (* constant 5138 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3227_t1 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) c) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a c) b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) c) (l_e_st_eq_landau_n_rt_rp_r_pl a (l_e_st_eq_landau_n_rt_rp_r_pl b c)) (l_e_st_eq_landau_n_rt_rp_r_pl a (l_e_st_eq_landau_n_rt_rp_r_pl c b)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a c) b) (l_e_st_eq_landau_n_rt_rp_r_asspl1 a b c) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_pl b c) (l_e_st_eq_landau_n_rt_rp_r_pl c b) a (l_e_st_eq_landau_n_rt_rp_r_compl b c)) (l_e_st_eq_landau_n_rt_rp_r_asspl2 a c b)))). Time Defined. (* constant 5139 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3227_t2 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) (l_e_st_eq_landau_n_rt_rp_r_pl c d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a c) (l_e_st_eq_landau_n_rt_rp_r_pl b d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) (l_e_st_eq_landau_n_rt_rp_r_pl c d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) c) d) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a c) b) d) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a c) (l_e_st_eq_landau_n_rt_rp_r_pl b d)) (l_e_st_eq_landau_n_rt_rp_r_asspl2 (l_e_st_eq_landau_n_rt_rp_r_pl a b) c d) (l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) c) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a c) b) d (l_e_st_eq_landau_n_rt_rp_r_c_3227_t1 a b c)) (l_e_st_eq_landau_n_rt_rp_r_asspl1 (l_e_st_eq_landau_n_rt_rp_r_pl a c) b d))))). Time Defined. (* constant 5140 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3227_t3 : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), (forall (c:l_e_st_eq_landau_n_rt_rp_r_real), (forall (d:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl a b) (l_e_st_eq_landau_n_rt_rp_r_pl c d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn a c) (l_e_st_eq_landau_n_rt_rp_r_mn b d)))))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (d:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl a b) (l_e_st_eq_landau_n_rt_rp_r_pl c d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl a b) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 c) (l_e_st_eq_landau_n_rt_rp_r_m0 d))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn a c) (l_e_st_eq_landau_n_rt_rp_r_mn b d)) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl c d)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 c) (l_e_st_eq_landau_n_rt_rp_r_m0 d)) (l_e_st_eq_landau_n_rt_rp_r_pl a b) (l_e_st_eq_landau_n_rt_rp_r_satz180 c d)) (l_e_st_eq_landau_n_rt_rp_r_c_3227_t2 a b (l_e_st_eq_landau_n_rt_rp_r_m0 c) (l_e_st_eq_landau_n_rt_rp_r_m0 d)))))). Time Defined. (* constant 5141 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3227_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_rets x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_rets x z)) (l_e_st_eq_landau_n_rt_rp_r_ismn12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_disttp2 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_disttp2 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_c_3227_t3 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))))). Time Defined. (* constant 5142 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3227_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x z)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_disttp2 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_disttp2 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_c_3227_t2 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im z)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))))). Time Defined. (* constant 5143 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3227_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_rets x z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x z))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_tsis2a x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_rets x z)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im z))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re z)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_3227_t4 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_3227_t5 x y z)) (l_e_st_eq_landau_n_rt_rp_r_c_plis12b (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x y) (l_e_st_eq_landau_n_rt_rp_r_c_rets x z) (l_e_st_eq_landau_n_rt_rp_r_c_imts x z))))). Time Defined. (* constant 5144 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz227 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_3227_t6 x y z))). Time Defined. (* constant 5145 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_disttp1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_comts (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_satz227 z x y) (l_e_st_eq_landau_n_rt_rp_r_c_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_comts z x) (l_e_st_eq_landau_n_rt_rp_r_c_comts z y))))). Time Defined. (* constant 5146 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_disttp2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz227 x y z))). Time Defined. (* constant 5147 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_distpt1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_disttp1 x y z)))). Time Defined. (* constant 5148 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_distpt2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_disttp2 x y z)))). Time Defined. (* constant 5149 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz228 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_m0 z))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 z))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_mn y z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_m0 z)) x (l_e_st_eq_landau_n_rt_rp_r_c_satz218 y z)) (l_e_st_eq_landau_n_rt_rp_r_c_disttp2 x y (l_e_st_eq_landau_n_rt_rp_r_c_m0 z)) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_m0 z)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz224b x z)) (l_e_st_eq_landau_n_rt_rp_r_c_satz218a (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))))). Time Defined. (* constant 5150 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_disttm1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_comts (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_satz228 z x y) (l_e_st_eq_landau_n_rt_rp_r_c_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_comts z x) (l_e_st_eq_landau_n_rt_rp_r_c_comts z y))))). Time Defined. (* constant 5151 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_disttm2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz228 x y z))). Time Defined. (* constant 5152 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_distmt1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_disttm1 x y z)))). Time Defined. (* constant 5153 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_distmt2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_disttm2 x y z)))). Time Defined. (* constant 5154 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (u1:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u2:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) x), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (u1:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u2:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) x) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x i j))))))). Time Defined. (* constant 5155 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (u1:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u2:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) x), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_mn u1 u2)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (u1:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u2:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) x) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_mn u1 u2)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_disttm2 y u1 u2) (l_e_st_eq_landau_n_rt_rp_r_c_satz213b (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t1 x y n u1 u2 i j))))))))). Time Defined. (* constant 5156 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (u1:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u2:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) x), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn u1 u2) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (u1:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u2:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) x) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x) => l_ore2 (l_e_st_eq_landau_n_rt_rp_r_c_is y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn u1 u2) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_satz221c y (l_e_st_eq_landau_n_rt_rp_r_c_mn u1 u2) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t2 x y n u1 u2 i j)) n))))))). Time Defined. (* constant 5157 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (u1:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u2:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) x), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x), l_e_st_eq_landau_n_rt_rp_r_c_is u1 u2))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (u1:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u2:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u1) x) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u2) x) => l_e_st_eq_landau_n_rt_rp_r_c_satz213a u1 u2 (l_e_st_eq_landau_n_rt_rp_r_c_3229_t3 x y n u1 u2 i j)))))))). Time Defined. (* constant 5158 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_pnot0 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_lemma4 y n)))). Time Defined. (* constant 5159 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_u : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_complex))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ov (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ov (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n)))) x))). Time Defined. (* constant 5160 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_dd : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_ov v (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n))))). Time Defined. (* constant 5161 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_satz197b (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_satz177 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_lemma6 (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n))))). Time Defined. (* constant 5162 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) l_e_st_eq_landau_n_rt_rp_r_1rl))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y)) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_lemma6 (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t5 x y n)) (l_e_st_eq_landau_n_rt_rp_r_lemma7 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n)) (l_e_st_eq_landau_n_rt_rp_r_lemma8 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n))))). Time Defined. (* constant 5163 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_satz197d (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_lemma6 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n))))). Time Defined. (* constant 5164 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t7 x y n) (l_e_st_eq_landau_n_rt_rp_r_lemma6 (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n))) (l_e_st_eq_landau_n_rt_rp_r_lemma7 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n))))). Time Defined. (* constant 5165 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_satz197a (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_satz182e (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))))). Time Defined. (* constant 5166 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_t8 x y n) (l_e_isf l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n t) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_t9 x y n)) (l_e_st_eq_landau_n_rt_rp_r_lemma9 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t4 x y n))))). Time Defined. (* constant 5167 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) l_e_st_eq_landau_n_rt_rp_r_c_1c))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y))))) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_tsis2a y (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im y) (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_t6 x y n) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t10 x y n))))). Time Defined. (* constant 5168 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_3229_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_3229_u x y n)) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_3229_u x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c x) x (l_e_st_eq_landau_n_rt_rp_r_c_assts2 y (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) x) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_3229_dd x y n (l_e_st_eq_landau_n_rt_rp_r_c_im y))))) l_e_st_eq_landau_n_rt_rp_r_c_1c x (l_e_st_eq_landau_n_rt_rp_r_c_3229_t11 x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_satz222b x)))). Time Defined. (* constant 5169 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_some (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y t) x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_somei l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y t) x) (l_e_st_eq_landau_n_rt_rp_r_c_3229_u x y n) (l_e_st_eq_landau_n_rt_rp_r_c_3229_t12 x y n)))). Time Defined. (* constant 5170 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_one (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y t) x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_onei l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y t) x) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y t) x) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) x) => l_e_st_eq_landau_n_rt_rp_r_c_satz229b x y n t u v w)))) (l_e_st_eq_landau_n_rt_rp_r_c_satz229a x y n)))). Time Defined. (* constant 5171 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ov : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_complex))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_ind l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y t) x) (l_e_st_eq_landau_n_rt_rp_r_c_satz229 x y n)))). Time Defined. (* constant 5172 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_oneax l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y t) x) (l_e_st_eq_landau_n_rt_rp_r_c_satz229 x y n)))). Time Defined. (* constant 5173 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229d : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) x (l_e_st_eq_landau_n_rt_rp_r_c_satz229c x y n)))). Time Defined. (* constant 5174 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229e : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) x (l_e_st_eq_landau_n_rt_rp_r_c_comts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) (l_e_st_eq_landau_n_rt_rp_r_c_satz229c x y n)))). Time Defined. (* constant 5175 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229f : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) x (l_e_st_eq_landau_n_rt_rp_r_c_satz229e x y n)))). Time Defined. (* constant 5176 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229g : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) x), l_e_st_eq_landau_n_rt_rp_r_c_is u (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) x) => l_e_st_eq_landau_n_rt_rp_r_c_satz229b x y n u (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) i (l_e_st_eq_landau_n_rt_rp_r_c_satz229c x y n)))))). Time Defined. (* constant 5177 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229h : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) u))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) x) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx u (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz229g x y u n i)))))). Time Defined. (* constant 5178 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229j : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts u y) x), l_e_st_eq_landau_n_rt_rp_r_c_is u (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts u y) x) => l_e_st_eq_landau_n_rt_rp_r_c_satz229g x y u n (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_ts u y) x (l_e_st_eq_landau_n_rt_rp_r_c_comts y u) i)))))). Time Defined. (* constant 5179 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz229k : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts u y) x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) u))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts u y) x) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx u (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz229j x y u n i)))))). Time Defined. (* constant 5180 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isov1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ov t z n) x y i))))). Time Defined. (* constant 5181 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isov2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov z x n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229h z x (l_e_st_eq_landau_n_rt_rp_r_c_ov z y o) n (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ov z y o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov z y o)) z (l_e_st_eq_landau_n_rt_rp_r_c_ists1 x y (l_e_st_eq_landau_n_rt_rp_r_c_ov z y o) i) (l_e_st_eq_landau_n_rt_rp_r_c_satz229c z y o)))))))). Time Defined. (* constant 5182 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isov12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z u), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y u o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z u) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y u o) (l_e_st_eq_landau_n_rt_rp_r_c_isov1 x y z i n) (l_e_st_eq_landau_n_rt_rp_r_c_isov2 z u y j n o))))))))). Time Defined. (* constant 5183 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz230 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) x (l_e_st_eq_landau_n_rt_rp_r_c_compl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) (l_e_st_eq_landau_n_rt_rp_r_c_satz212h x y))). Time Defined. (* constant 5184 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz231 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) y) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212e (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) y x (l_e_st_eq_landau_n_rt_rp_r_c_compl y x))). Time Defined. (* constant 5185 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz232 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) y)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212e x (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y (l_e_st_eq_landau_n_rt_rp_r_c_satz230 x y))). Time Defined. (* constant 5186 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4233_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) z) y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl z y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) z) y) (l_e_st_eq_landau_n_rt_rp_r_c_compl (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z)) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_pl z y) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_compl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_asspl2 (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) z y)))). Time Defined. (* constant 5187 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4233_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z)) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) z) y) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) x (l_e_st_eq_landau_n_rt_rp_r_c_4233_t1 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y (l_e_st_eq_landau_n_rt_rp_r_c_satz230 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z)) (l_e_st_eq_landau_n_rt_rp_r_c_satz230 x y)))). Time Defined. (* constant 5188 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz233 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212d x (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_4233_t2 x y z)))). Time Defined. (* constant 5189 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz234 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212g (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn y z) z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) (l_e_st_eq_landau_n_rt_rp_r_c_asspl1 x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z) z) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn y z) z) y x (l_e_st_eq_landau_n_rt_rp_r_c_satz230 y z)))))). Time Defined. (* constant 5190 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz234a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_satz234 x y z)))). Time Defined. (* constant 5191 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz234b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl y x) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_mn x z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) y) (l_e_st_eq_landau_n_rt_rp_r_c_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y x) z (l_e_st_eq_landau_n_rt_rp_r_c_compl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz234 y x z) (l_e_st_eq_landau_n_rt_rp_r_c_compl y (l_e_st_eq_landau_n_rt_rp_r_c_mn x z))))). Time Defined. (* constant 5192 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz234c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) y) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) y) (l_e_st_eq_landau_n_rt_rp_r_c_satz234b x y z)))). Time Defined. (* constant 5193 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz235 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212f x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl z (l_e_st_eq_landau_n_rt_rp_r_c_mn y z))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) x (l_e_st_eq_landau_n_rt_rp_r_c_asspl1 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_pl z (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) y (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz212h y z)) (l_e_st_eq_landau_n_rt_rp_r_c_satz230 x y))))). Time Defined. (* constant 5194 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz235a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_satz235 x y z)))). Time Defined. (* constant 5195 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz235b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) y) (l_e_st_eq_landau_n_rt_rp_r_c_satz235a x y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz234c x z y)))). Time Defined. (* constant 5196 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz235c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl z x) y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn x (l_e_st_eq_landau_n_rt_rp_r_c_mn y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) y) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl z x) y) (l_e_st_eq_landau_n_rt_rp_r_c_satz235b x y z) (l_e_st_eq_landau_n_rt_rp_r_c_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl z x) y (l_e_st_eq_landau_n_rt_rp_r_c_compl x z))))). Time Defined. (* constant 5197 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz236 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212g (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_asspl2 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y z) (l_e_st_eq_landau_n_rt_rp_r_c_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) x z (l_e_st_eq_landau_n_rt_rp_r_c_satz230 x y)))))). Time Defined. (* constant 5198 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz236a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl z x) (l_e_st_eq_landau_n_rt_rp_r_c_pl z y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl z x) (l_e_st_eq_landau_n_rt_rp_r_c_pl z y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_pl z x) (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl z y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_compl z x) (l_e_st_eq_landau_n_rt_rp_r_c_compl z y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz236 x y z)))). Time Defined. (* constant 5199 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4237_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) (l_e_st_eq_landau_n_rt_rp_r_c_pl u y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) (l_e_st_eq_landau_n_rt_rp_r_c_pl u y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) u) y) (l_e_st_eq_landau_n_rt_rp_r_c_pl z y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_asspl2 (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) u y) (l_e_st_eq_landau_n_rt_rp_r_c_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) u) z y (l_e_st_eq_landau_n_rt_rp_r_c_satz230 z u)) (l_e_st_eq_landau_n_rt_rp_r_c_compl z y))))). Time Defined. (* constant 5200 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4237_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl y u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl y u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl u y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) (l_e_st_eq_landau_n_rt_rp_r_c_pl u y))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) (l_e_st_eq_landau_n_rt_rp_r_c_pl u y) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_compl y u)) (l_e_st_eq_landau_n_rt_rp_r_c_asspl1 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) (l_e_st_eq_landau_n_rt_rp_r_c_pl u y)) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) (l_e_st_eq_landau_n_rt_rp_r_c_pl u y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_4237_t1 x y z u)))))). Time Defined. (* constant 5201 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4237_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl y u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl y u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_4237_t2 x y z u) (l_e_st_eq_landau_n_rt_rp_r_c_asspl2 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y z) (l_e_st_eq_landau_n_rt_rp_r_c_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) x z (l_e_st_eq_landau_n_rt_rp_r_c_satz230 x y)))))). Time Defined. (* constant 5202 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz237 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y u)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212f (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_pl y u) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_4237_t3 x y z u))))). Time Defined. (* constant 5203 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4238_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl z u)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl u z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl z u)) (l_e_st_eq_landau_n_rt_rp_r_c_asspl1 x u z) (l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_pl u z) (l_e_st_eq_landau_n_rt_rp_r_c_pl z u) x (l_e_st_eq_landau_n_rt_rp_r_c_compl u z)))))). Time Defined. (* constant 5204 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4238_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) u)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl z u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_pl z u))) (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz237 (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) z u) (l_e_st_eq_landau_n_rt_rp_r_c_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl x (l_e_st_eq_landau_n_rt_rp_r_c_pl z u)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_pl z u)) (l_e_st_eq_landau_n_rt_rp_r_c_4238_t1 x y z u) (l_e_st_eq_landau_n_rt_rp_r_c_asspl1 y z u)) (l_e_st_eq_landau_n_rt_rp_r_c_satz236 x y (l_e_st_eq_landau_n_rt_rp_r_c_pl z u)))))). Time Defined. (* constant 5205 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz238 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212g (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) (l_e_st_eq_landau_n_rt_rp_r_c_4238_t2 x y z u))))). Time Defined. (* constant 5206 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4239_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_satz238 x y z u) (l_e_st_eq_landau_n_rt_rp_r_c_satz213b (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) i)))))). Time Defined. (* constant 5207 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz239a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) => l_e_st_eq_landau_n_rt_rp_r_c_satz213a (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) (l_e_st_eq_landau_n_rt_rp_r_c_4239_t1 x y z u i)))))). Time Defined. (* constant 5208 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_4239_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_satz238 x y z u) (l_e_st_eq_landau_n_rt_rp_r_c_satz213b (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z) i)))))). Time Defined. (* constant 5209 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz239b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x u) (l_e_st_eq_landau_n_rt_rp_r_c_pl y z)) => l_e_st_eq_landau_n_rt_rp_r_c_satz213a (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) (l_e_st_eq_landau_n_rt_rp_r_c_mn z u) (l_e_st_eq_landau_n_rt_rp_r_c_4239_t2 x y z u i)))))). Time Defined. (* constant 5210 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz240 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) x (l_e_st_eq_landau_n_rt_rp_r_c_comts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) (l_e_st_eq_landau_n_rt_rp_r_c_satz229c x y n)))). Time Defined. (* constant 5211 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz241 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) y n) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229h (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) y x n (l_e_st_eq_landau_n_rt_rp_r_c_comts y x)))). Time Defined. (* constant 5212 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ov x y o) l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y o) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) n (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y o) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx x l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y o)) (l_e_st_eq_landau_n_rt_rp_r_c_satz229c x y o) (l_e_st_eq_landau_n_rt_rp_r_c_satz221b y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y o) t)))))). Time Defined. (* constant 5213 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz242 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ov x y o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 x y n o)) y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229h x (l_e_st_eq_landau_n_rt_rp_r_c_ov x y o) y (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 x y n o) (l_e_st_eq_landau_n_rt_rp_r_c_satz240 x y o))))). Time Defined. (* constant 5214 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5243_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) z) y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) z) y) (l_e_st_eq_landau_n_rt_rp_r_c_comts (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o)) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) (l_e_st_eq_landau_n_rt_rp_r_c_comts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_assts2 (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) z y)))))). Time Defined. (* constant 5215 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5243_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o)) x))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) z) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) x (l_e_st_eq_landau_n_rt_rp_r_c_5243_t1 x y z n o) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) z) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y (l_e_st_eq_landau_n_rt_rp_r_c_satz240 (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o)) (l_e_st_eq_landau_n_rt_rp_r_c_satz240 x y n)))))). Time Defined. (* constant 5216 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz243 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229g x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z o) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o) (l_e_st_eq_landau_n_rt_rp_r_c_5243_t2 x y z n o)))))). Time Defined. (* constant 5217 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz244 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229k (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)) n (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n) z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n) z) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n) z) y x (l_e_st_eq_landau_n_rt_rp_r_c_satz240 y z n))))))). Time Defined. (* constant 5218 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz244a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)) (l_e_st_eq_landau_n_rt_rp_r_c_satz244 x y z n))))). Time Defined. (* constant 5219 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz244b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts y x) z n) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) y) (l_e_st_eq_landau_n_rt_rp_r_c_isov1 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y x) z (l_e_st_eq_landau_n_rt_rp_r_c_comts x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_satz244 y x z n) (l_e_st_eq_landau_n_rt_rp_r_c_comts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n)))))). Time Defined. (* constant 5220 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz244c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) y) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) y) (l_e_st_eq_landau_n_rt_rp_r_c_satz244b x y z n))))). Time Defined. (* constant 5221 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz245 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z) (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 y z n o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229j x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 y z n o) (l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) x (l_e_st_eq_landau_n_rt_rp_r_c_assts1 (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o)) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o)) y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz229c y z o)) (l_e_st_eq_landau_n_rt_rp_r_c_satz240 x y n))))))). Time Defined. (* constant 5222 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz245a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z) (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_satz245 x y z n o)))))). Time Defined. (* constant 5223 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz245b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) y n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) z) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz245a x y z n o) (l_e_st_eq_landau_n_rt_rp_r_c_satz244c x z y n)))))). Time Defined. (* constant 5224 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz245c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) y n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_ov y z o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz245b x y z n o) (l_e_st_eq_landau_n_rt_rp_r_c_isov1 (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) y (l_e_st_eq_landau_n_rt_rp_r_c_comts x z) n)))))). Time Defined. (* constant 5225 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz246 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229k (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_assts2 (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y z) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) x z (l_e_st_eq_landau_n_rt_rp_r_c_satz240 x y n)))))))). Time Defined. (* constant 5226 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz246a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d z y o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d z y o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_isov12 (l_e_st_eq_landau_n_rt_rp_r_c_ts z x) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_comts z x) (l_e_st_eq_landau_n_rt_rp_r_c_comts z y) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d z y o n) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_satz246 x y z n o)))))). Time Defined. (* constant 5227 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5247_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_ts u y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_ts u y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) u) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts z y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_assts2 (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) u y) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) u) z y (l_e_st_eq_landau_n_rt_rp_r_c_satz240 z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_comts z y))))))). Time Defined. (* constant 5228 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5247_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts u y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_ts u y))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_ts u y) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_comts y u)) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_ts u y)) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_ts u y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_5247_t1 x y z u n o)))))))). Time Defined. (* constant 5229 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5247_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_5247_t2 x y z u n o) (l_e_st_eq_landau_n_rt_rp_r_c_assts2 (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y z) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) x z (l_e_st_eq_landau_n_rt_rp_r_c_satz240 x y n)))))))). Time Defined. (* constant 5230 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz247 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229j (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o) (l_e_st_eq_landau_n_rt_rp_r_c_5247_t3 x y z u n o))))))). Time Defined. (* constant 5231 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5248_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts z u))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts u z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 x u z) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ts u z) (l_e_st_eq_landau_n_rt_rp_r_c_ts z u) x (l_e_st_eq_landau_n_rt_rp_r_c_comts u z))))))))). Time Defined. (* constant 5232 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5248_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u p)) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u p)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) u (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o) p)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y (l_e_st_eq_landau_n_rt_rp_r_c_ts z u) n (l_e_st_eq_landau_n_rt_rp_r_c_satz221d z u o p))) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz247 (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) z u (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o) p) (l_e_st_eq_landau_n_rt_rp_r_c_isov12 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_5248_t1 x y z u n o p) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 y z u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) u (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o) p) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y (l_e_st_eq_landau_n_rt_rp_r_c_ts z u) n (l_e_st_eq_landau_n_rt_rp_r_c_satz221d z u o p))) (l_e_st_eq_landau_n_rt_rp_r_c_satz246 x y (l_e_st_eq_landau_n_rt_rp_r_c_ts z u) n (l_e_st_eq_landau_n_rt_rp_r_c_satz221d z u o p))))))))). Time Defined. (* constant 5233 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz248 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u p) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 z u o p)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229k (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u p) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n o)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 z u o p) (l_e_st_eq_landau_n_rt_rp_r_c_5248_t2 x y z u n o p)))))))). Time Defined. (* constant 5234 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz249 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_0c x n) l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229h l_e_st_eq_landau_n_rt_rp_r_c_0c x l_e_st_eq_landau_n_rt_rp_r_c_0c n (l_e_st_eq_landau_n_rt_rp_r_c_satz221b x l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_0c)))). Time Defined. (* constant 5235 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz250 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x x n) l_e_st_eq_landau_n_rt_rp_r_c_1c)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229h x x l_e_st_eq_landau_n_rt_rp_r_c_1c n (l_e_st_eq_landau_n_rt_rp_r_c_satz222 x))). Time Defined. (* constant 5236 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz251a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) l_e_st_eq_landau_n_rt_rp_r_c_1c)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov x x (l_e_notis_th2 l_e_st_eq_landau_n_rt_rp_r_c_cx y l_e_st_eq_landau_n_rt_rp_r_c_0c x n i)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_isov2 y x x (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx x y i) n (l_e_notis_th2 l_e_st_eq_landau_n_rt_rp_r_c_cx y l_e_st_eq_landau_n_rt_rp_r_c_0c x n i)) (l_e_st_eq_landau_n_rt_rp_r_c_satz250 x (l_e_notis_th2 l_e_st_eq_landau_n_rt_rp_r_c_cx y l_e_st_eq_landau_n_rt_rp_r_c_0c x n i)))))). Time Defined. (* constant 5237 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz251b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) l_e_st_eq_landau_n_rt_rp_r_c_1c), l_e_st_eq_landau_n_rt_rp_r_c_is x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) l_e_st_eq_landau_n_rt_rp_r_c_1c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y l_e_st_eq_landau_n_rt_rp_r_c_1c) y (l_e_st_eq_landau_n_rt_rp_r_c_satz229d x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) l_e_st_eq_landau_n_rt_rp_r_c_1c y i) (l_e_st_eq_landau_n_rt_rp_r_c_satz222 y))))). Time Defined. (* constant 5238 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_0c u o) l_e_st_eq_landau_n_rt_rp_r_c_0c i (l_e_st_eq_landau_n_rt_rp_r_c_isov1 z l_e_st_eq_landau_n_rt_rp_r_c_0c u j o) (l_e_st_eq_landau_n_rt_rp_r_c_satz249 u o))))))))). Time Defined. (* constant 5239 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_satz229d x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz221b y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_5252_t1 x y z u n o i j)))))))))). Time Defined. (* constant 5240 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_satz221a x u (l_e_st_eq_landau_n_rt_rp_r_c_5252_t2 x y z u n o i j)) (l_e_st_eq_landau_n_rt_rp_r_c_satz221b y z j))))))))). Time Defined. (* constant 5241 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n p)) l_e_st_eq_landau_n_rt_rp_r_c_1c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n p)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 z u p o)) (l_e_st_eq_landau_n_rt_rp_r_c_satz248 x y z u n p o) (l_e_st_eq_landau_n_rt_rp_r_c_satz251a (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 z u p o) i))))))))). Time Defined. (* constant 5242 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz251b (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n p) (l_e_st_eq_landau_n_rt_rp_r_c_5252_t4 x y z u n o i p))))))))). Time Defined. (* constant 5243 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz252a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_5252_t3 x y z u n o i t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_5252_t5 x y z u n o i t)))))))). Time Defined. (* constant 5244 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) l_e_st_eq_landau_n_rt_rp_r_c_0c i (l_e_st_eq_landau_n_rt_rp_r_c_satz221b y z j))))))))). Time Defined. (* constant 5245 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is u l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_satz221c x u (l_e_st_eq_landau_n_rt_rp_r_c_5252_t6 x y z u n o i j)) o)))))))). Time Defined. (* constant 5246 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_0c y n) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_isov1 x l_e_st_eq_landau_n_rt_rp_r_c_0c y (l_e_st_eq_landau_n_rt_rp_r_c_5252_t7 x y z u n o i j) n) (l_e_st_eq_landau_n_rt_rp_r_c_satz249 y n)) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_0c u o) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_isov1 z l_e_st_eq_landau_n_rt_rp_r_c_0c u j o) (l_e_st_eq_landau_n_rt_rp_r_c_satz249 u o)))))))))). Time Defined. (* constant 5247 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 z u p o)) l_e_st_eq_landau_n_rt_rp_r_c_1c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 z u p o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n p)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_satz248 x y z u n p o) (l_e_st_eq_landau_n_rt_rp_r_c_satz251a (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y z n p) i))))))))). Time Defined. (* constant 5248 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_5252_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz251b (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma6 z u p o) (l_e_st_eq_landau_n_rt_rp_r_c_5252_t9 x y z u n o i p))))))))). Time Defined. (* constant 5249 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz252b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_5252_t8 x y z u n o i t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_5252_t10 x y z u n o i t)))))))). Time Defined. (* constant 5250 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz253 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) y n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229g (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) y (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n)) n (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n))) (l_e_st_eq_landau_n_rt_rp_r_c_pl x z) (l_e_st_eq_landau_n_rt_rp_r_c_disttp2 y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) x (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n)) z (l_e_st_eq_landau_n_rt_rp_r_c_satz229c x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz229c z y n))))))). Time Defined. (* constant 5251 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_distop : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_satz253 x z y n))))). Time Defined. (* constant 5252 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_distpo : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) z n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz253 x z y n)))). Time Defined. (* constant 5253 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz254 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o))) (l_e_st_eq_landau_n_rt_rp_r_c_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_satz246 x y u n o) (l_e_st_eq_landau_n_rt_rp_r_c_satz246a z u y o n)) (l_e_st_eq_landau_n_rt_rp_r_c_satz253 (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)))))))). Time Defined. (* constant 5254 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz255 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) y n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229g (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) y (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n)) n (l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n))) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n))) (l_e_st_eq_landau_n_rt_rp_r_c_mn x z) (l_e_st_eq_landau_n_rt_rp_r_c_disttm2 y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) x (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ov z y n)) z (l_e_st_eq_landau_n_rt_rp_r_c_satz229c x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz229c z y n))))))). Time Defined. (* constant 5255 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_distom : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z n) (l_e_st_eq_landau_n_rt_rp_r_c_satz255 x z y n))))). Time Defined. (* constant 5256 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_distmo : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov x z n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y z n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) z n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis z l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz255 x z y n)))). Time Defined. (* constant 5257 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz256 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis u l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o))) (l_e_st_eq_landau_n_rt_rp_r_c_ismn12 (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov z u o) (l_e_st_eq_landau_n_rt_rp_r_c_satz246 x y u n o) (l_e_st_eq_landau_n_rt_rp_r_c_satz246a z u y o n)) (l_e_st_eq_landau_n_rt_rp_r_c_satz255 (l_e_st_eq_landau_n_rt_rp_r_c_ts x u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u) (l_e_st_eq_landau_n_rt_rp_r_c_ts y z) (l_e_st_eq_landau_n_rt_rp_r_c_satz221d y u n o)))))))). Time Defined. (* constant 5258 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_conj : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_complex). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))). Time Defined. (* constant 5259 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_conjisa : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_pli a (l_e_st_eq_landau_n_rt_rp_r_m0 b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) a (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b))) (l_e_st_eq_landau_n_rt_rp_r_m0 b) (l_e_st_eq_landau_n_rt_rp_r_c_reis a b) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) b (l_e_st_eq_landau_n_rt_rp_r_c_imis a b)))). Time Defined. (* constant 5260 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_conjisb : (forall (a:l_e_st_eq_landau_n_rt_rp_r_real), (forall (b:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli a (l_e_st_eq_landau_n_rt_rp_r_m0 b)) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)))). exact (fun (a:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (b:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_pli a b)) (l_e_st_eq_landau_n_rt_rp_r_c_pli a (l_e_st_eq_landau_n_rt_rp_r_m0 b)) (l_e_st_eq_landau_n_rt_rp_r_c_conjisa a b))). Time Defined. (* constant 5261 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isconj : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_conj t) x y i))). Time Defined. (* constant 5262 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz257 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) x (l_e_st_eq_landau_n_rt_rp_r_c_conjisa (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_satz177 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_pliis x)). Time Defined. (* constant 5263 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz258a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_isconj x l_e_st_eq_landau_n_rt_rp_r_c_0c i) (l_e_st_eq_landau_n_rt_rp_r_c_conjisa l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satz176b l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 5264 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6258_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_isre (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_lemma1 (l_e_st_eq_landau_n_rt_rp_r_c_conj x) i))). Time Defined. (* constant 5265 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6258_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_satz176e (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_isim (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_lemma2 (l_e_st_eq_landau_n_rt_rp_r_c_conj x) i)))). Time Defined. (* constant 5266 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz258b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_ispli x) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_re x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_6258_t1 x i) (l_e_st_eq_landau_n_rt_rp_r_c_6258_t2 x i)))). Time Defined. (* constant 5267 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6258_anders : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx x l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_satz257 x) (l_e_st_eq_landau_n_rt_rp_r_c_satz258a (l_e_st_eq_landau_n_rt_rp_r_c_conj x) i))). Time Defined. (* constant 5268 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz258c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) n (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz258b x t))). Time Defined. (* constant 5269 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6259_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) x) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_isim (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_isceim (l_e_st_eq_landau_n_rt_rp_r_c_conj x) x i))). Time Defined. (* constant 5270 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz259a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) x) => l_e_st_eq_landau_n_rt_rp_r_lemma10 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_6259_t1 x i)))). Time Defined. (* constant 5271 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz259b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) x (l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satz176b (l_e_st_eq_landau_n_rt_rp_r_c_im x) i) i)) (l_e_st_eq_landau_n_rt_rp_r_c_pliis x))). Time Defined. (* constant 5272 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz269c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) => l_e_st_eq_landau_n_rt_rp_r_c_satz259a x (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_conj x) i))). Time Defined. (* constant 5273 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz269d : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_conj x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj x) x (l_e_st_eq_landau_n_rt_rp_r_c_satz259b x i))). Time Defined. (* constant 5274 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz260 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_conjisa (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_satz180 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_plis12b (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))))). Time Defined. (* constant 5275 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz260a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz260 x y))). Time Defined. (* constant 5276 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6261_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_imts x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_imts x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_satz180 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_satz197f (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_satz197e (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re y))))). Time Defined. (* constant 5277 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz261 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_imts x y))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y)))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_conjisa (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_c_imts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_c_rets x y) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y)))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_imts x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y))) (l_e_st_eq_landau_n_rt_rp_r_ismn2 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_satz198a (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_6261_t1 x y)) (l_e_st_eq_landau_n_rt_rp_r_c_tsis12b (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im y))))). Time Defined. (* constant 5278 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz261a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz261 x y))). Time Defined. (* constant 5279 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6262_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) x (l_e_st_eq_landau_n_rt_rp_r_c_satz230 x y))). Time Defined. (* constant 5280 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6262_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_isconj x (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y) (l_e_st_eq_landau_n_rt_rp_r_c_6262_t1 x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz260 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y) y))). Time Defined. (* constant 5281 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz262 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz212f (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_6262_t2 x y)))). Time Defined. (* constant 5282 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz262a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz262 x y))). Time Defined. (* constant 5283 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6263_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229f x y n))). Time Defined. (* constant 5284 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6263_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_isconj x (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) (l_e_st_eq_landau_n_rt_rp_r_c_6263_t1 x y n)))). Time Defined. (* constant 5285 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6263_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_conj y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz261 (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y))). Time Defined. (* constant 5286 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6263_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_conj y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_6263_t2 x y n) (l_e_st_eq_landau_n_rt_rp_r_c_6263_t3 x y n)))). Time Defined. (* constant 5287 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_6263_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_conj y) l_e_st_eq_landau_n_rt_rp_r_c_0c))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz258c y n))). Time Defined. (* constant 5288 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz263 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y) (l_e_st_eq_landau_n_rt_rp_r_c_satz258c y n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229j (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_6263_t5 x y n) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_6263_t4 x y n))))). Time Defined. (* constant 5289 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz263a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y) (l_e_st_eq_landau_n_rt_rp_r_c_satz258c y n)) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y) (l_e_st_eq_landau_n_rt_rp_r_c_satz258c y n)) (l_e_st_eq_landau_n_rt_rp_r_c_satz263 x y n)))). Time Defined. (* constant 5290 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_mod : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_sqrt (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 x)). Time Defined. (* constant 5291 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ismod : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_mod t) x y i))). Time Defined. (* constant 5292 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz264a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_sqrtnot0 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 x) (l_e_st_eq_landau_n_rt_rp_r_pnot0 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma4 x n)))). Time Defined. (* constant 5293 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz264b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_sqrt0 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma3 x i))). Time Defined. (* constant 5294 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz264c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_thsqrt1a (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 x)). Time Defined. (* constant 5295 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz264d : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_satz167f (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0 (l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_c_satz264c x) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz169d (l_e_st_eq_landau_n_rt_rp_r_c_mod x) t))). Time Defined. (* constant 5296 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_thsqrt1b (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 x))). Time Defined. (* constant 5297 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_lemma12 (l_e_st_eq_landau_n_rt_rp_r_c_re x))). Time Defined. (* constant 5298 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_satz191 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_moreisi2 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)))) (l_e_st_eq_landau_n_rt_rp_r_lemma11 (l_e_st_eq_landau_n_rt_rp_r_c_im x))). Time Defined. (* constant 5299 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ismoreis12 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t1 x) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t2 x) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t3 x)). Time Defined. (* constant 5300 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_lemma12 (l_e_st_eq_landau_n_rt_rp_r_c_im x))). Time Defined. (* constant 5301 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_satz191 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_lemma11 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_moreisi2 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))))). Time Defined. (* constant 5302 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ismoreis12 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t1 x) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t5 x) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t6 x)). Time Defined. (* constant 5303 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_more s r))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_lemma2 r s l))))). Time Defined. (* constant 5304 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_pos s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_satz169b s (l_e_st_eq_landau_n_rt_rp_r_satz172d s r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_7265_t8 r s m n l) n)))))). Time Defined. (* constant 5305 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (o:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_ts r r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_trmore (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_satz203a s r s (l_e_st_eq_landau_n_rt_rp_r_c_7265_t8 r s m n l) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t9 r s m n l)) (l_e_st_eq_landau_n_rt_rp_r_satz203d s r r (l_e_st_eq_landau_n_rt_rp_r_c_7265_t8 r s m n l) (l_e_st_eq_landau_n_rt_rp_r_satz169b r o)))))))). Time Defined. (* constant 5306 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_ts r r))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ismore2 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 r r i)) (l_e_st_eq_landau_n_rt_rp_r_satz169a (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_possq s (l_e_st_eq_landau_n_rt_rp_r_pnot0 s (l_e_st_eq_landau_n_rt_rp_r_c_7265_t9 r s m n l)))))))))). Time Defined. (* constant 5307 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0), (forall (l:l_e_st_eq_landau_n_rt_rp_r_less r s), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_lemma1 (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_orapp (l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_ts r r)) n (fun (t:l_e_st_eq_landau_n_rt_rp_r_more r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_7265_t10 r s m n l t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_7265_t11 r s m n l t))))))). Time Defined. (* constant 5308 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7265_t13 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_moreis r s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_satz167f r s (fun (t:l_e_st_eq_landau_n_rt_rp_r_less r s) => l_e_st_eq_landau_n_rt_rp_r_satz167c (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s) m (l_e_st_eq_landau_n_rt_rp_r_c_7265_t12 r s m n t)))))). Time Defined. (* constant 5309 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz265a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_7265_t13 (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t4 x) (l_e_st_eq_landau_n_rt_rp_r_c_satz264d x)). Time Defined. (* constant 5310 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz265b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_7265_t13 (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_7265_t7 x) (l_e_st_eq_landau_n_rt_rp_r_c_satz264d x)). Time Defined. (* constant 5311 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7266_t1 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts t t) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts t t)). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts t t) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_satz176b (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 5312 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7266_t2 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts t l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 t)) l_e_st_eq_landau_n_rt_rp_r_0). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts t l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 t)) (l_e_st_eq_landau_n_rt_rp_r_ts t l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts t l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 t) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 t (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ts02 t l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 5313 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7266_t3 : (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli t l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli t l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts t t) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli t l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli t l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts t t) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts t l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 t))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts t t) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_tsis12a t l_e_st_eq_landau_n_rt_rp_r_0 t l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts t t) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts t t) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts t l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 t)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_7266_t1 t) (l_e_st_eq_landau_n_rt_rp_r_c_7266_t2 t))). Time Defined. (* constant 5314 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7266_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts s s) l_e_st_eq_landau_n_rt_rp_r_0)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts s s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_7266_t3 r)) i (l_e_st_eq_landau_n_rt_rp_r_c_7266_t3 s)))))). Time Defined. (* constant 5315 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7266_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts s s) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_c_isre (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_iscere (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts s s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_7266_t4 r s i n o)) (l_e_st_eq_landau_n_rt_rp_r_c_reis (l_e_st_eq_landau_n_rt_rp_r_ts s s) l_e_st_eq_landau_n_rt_rp_r_0)))))). Time Defined. (* constant 5316 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7266_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)), l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) => l_andi (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts s s)) n (l_e_st_eq_landau_n_rt_rp_r_c_7266_t5 r s i n o)))))). Time Defined. (* constant 5317 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7266_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)), l_and (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_ts s s))))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) => l_andi (l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts s s) (l_e_st_eq_landau_n_rt_rp_r_ts s s)) o (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts s s))))))). Time Defined. (* constant 5318 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz266 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)), l_e_st_eq_landau_n_rt_rp_r_is r s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0))) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_neg s)) => l_e_st_eq_landau_n_rt_rp_r_satzr161b (l_e_st_eq_landau_n_rt_rp_r_ts s s) r s (l_e_st_eq_landau_n_rt_rp_r_c_7266_t6 r s i n o) (l_e_st_eq_landau_n_rt_rp_r_c_7266_t7 r s i n o)))))). Time Defined. (* constant 5319 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7267_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_7266_t3 (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx1 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_thsqrt1b (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 x)))). Time Defined. (* constant 5320 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7267_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_satz197e (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_satz198 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). Time Defined. (* constant 5321 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7267_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re x))) l_e_st_eq_landau_n_rt_rp_r_0). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_satz197b (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_satz179a (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). Time Defined. (* constant 5322 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7267_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_tsis2a x (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_re x))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_7267_t2 x) (l_e_st_eq_landau_n_rt_rp_r_c_7267_t3 x))). Time Defined. (* constant 5323 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz267 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_7267_t1 x) (l_e_st_eq_landau_n_rt_rp_r_c_7267_t4 x)). Time Defined. (* constant 5324 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz267a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_satz267 x)). Time Defined. (* constant 5325 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ts x z))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts y x) z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ts x z)) (l_e_st_eq_landau_n_rt_rp_r_c_assts2 x y z) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts y x) z (l_e_st_eq_landau_n_rt_rp_r_c_comts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 y x z)))). Time Defined. (* constant 5326 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ts z u))) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_ts y u))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x z) (l_e_st_eq_landau_n_rt_rp_r_c_ts y u)) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 x y (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_ts z u)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_ts y u)) x (l_e_st_eq_landau_n_rt_rp_r_c_7268_t1 y z u)) (l_e_st_eq_landau_n_rt_rp_r_c_assts2 x z (l_e_st_eq_landau_n_rt_rp_r_c_ts y u)))))). Time Defined. (* constant 5327 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_conj y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts x y))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_conj y))) (l_e_st_eq_landau_n_rt_rp_r_c_satz267 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_conj (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz261 x y)) (l_e_st_eq_landau_n_rt_rp_r_c_7268_t2 x y (l_e_st_eq_landau_n_rt_rp_r_c_conj x) (l_e_st_eq_landau_n_rt_rp_r_c_conj y)))). Time Defined. (* constant 5328 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_conj y))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_7268_t3 x y) (l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_conj x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts y (l_e_st_eq_landau_n_rt_rp_r_c_conj y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_satz267a x) (l_e_st_eq_landau_n_rt_rp_r_c_satz267a y)) (l_e_st_eq_landau_n_rt_rp_r_c_7268_t2 (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 5329 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_satz176b (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 5330 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s)) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s)) (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 s (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ts02 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 5331 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_tsis12a r l_e_st_eq_landau_n_rt_rp_r_0 s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_7268_t5 r s) (l_e_st_eq_landau_n_rt_rp_r_c_7268_t6 r s)))). Time Defined. (* constant 5332 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_7268_t4 x y) (l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_7268_t7 (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_7268_t7 (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))))). Time Defined. (* constant 5333 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7268_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))), l_con))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) => l_orapp (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) (l_and (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) l_con (l_e_st_eq_landau_n_rt_rp_r_satz196h (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y) n) (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) => l_e_st_eq_landau_n_rt_rp_r_c_satz264c y (l_ande2 (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) t)) (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) => l_e_st_eq_landau_n_rt_rp_r_c_satz264c x (l_ande1 (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) t))))). Time Defined. (* constant 5334 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz268 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz266 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_7268_t8 x y) (l_e_st_eq_landau_n_rt_rp_r_c_satz264c (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) => l_e_st_eq_landau_n_rt_rp_r_c_7268_t9 x y t))). Time Defined. (* constant 5335 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz268a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz268 x y))). Time Defined. (* constant 5336 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7269_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_pnot0 (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_satz264a y n)))). Time Defined. (* constant 5337 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7269_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz268 (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) (l_e_st_eq_landau_n_rt_rp_r_c_ismod (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n) y) x (l_e_st_eq_landau_n_rt_rp_r_c_satz240 x y n))))). Time Defined. (* constant 5338 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7269_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_ov (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_7269_t1 x y n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_satz204g (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_7269_t1 x y n) (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_comts (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n))) (l_e_st_eq_landau_n_rt_rp_r_c_7269_t2 x y n))))). Time Defined. (* constant 5339 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz269 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x y n)) (l_e_st_eq_landau_n_rt_rp_r_ov (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_pnot0 (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_satz264a y n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_7269_t3 x y n))). Time Defined. (* constant 5340 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7270_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_abs r) r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_abs r) r) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_moreisi1 (l_e_st_eq_landau_n_rt_rp_r_abs r) r (l_e_st_eq_landau_n_rt_rp_r_trmore (l_e_st_eq_landau_n_rt_rp_r_abs r) l_e_st_eq_landau_n_rt_rp_r_0 r (l_e_st_eq_landau_n_rt_rp_r_satz169a (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_satz166b r t)) (l_e_st_eq_landau_n_rt_rp_r_lemma2 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satz169c r t)))) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => l_e_st_eq_landau_n_rt_rp_r_moreisi2 (l_e_st_eq_landau_n_rt_rp_r_abs r) r (l_e_st_eq_landau_n_rt_rp_r_absnn r t))). Time Defined. (* constant 5341 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7270_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_trmoreis (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_satz265a x) (l_e_st_eq_landau_n_rt_rp_r_c_7270_t1 (l_e_st_eq_landau_n_rt_rp_r_c_re x))). Time Defined. (* constant 5342 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7270_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_1c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_1rl))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_1c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_re l_e_st_eq_landau_n_rt_rp_r_c_1c) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_isre (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im y))) (l_e_st_eq_landau_n_rt_rp_r_c_iscere (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_1c i) (l_e_st_eq_landau_n_rt_rp_r_c_reis l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 5343 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz270 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_1c), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) l_e_st_eq_landau_n_rt_rp_r_1rl))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_1c) => l_e_st_eq_landau_n_rt_rp_r_ismoreis2 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re y)) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_7270_t3 x y i) (l_e_st_eq_landau_n_rt_rp_r_satz191 (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_re y) (l_e_st_eq_landau_n_rt_rp_r_c_7270_t2 x) (l_e_st_eq_landau_n_rt_rp_r_c_7270_t2 y))))). Time Defined. (* constant 5344 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz264b (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) i))). Time Defined. (* constant 5345 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_ismoreis2 (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t1 x y i)) (l_e_st_eq_landau_n_rt_rp_r_satz191 (l_e_st_eq_landau_n_rt_rp_r_c_mod x) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_mod y) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_satz264d x) (l_e_st_eq_landau_n_rt_rp_r_c_satz264d y))))). Time Defined. (* constant 5346 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_nis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_pnot0 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz264a (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)))). Time Defined. (* constant 5347 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)) l_e_st_eq_landau_n_rt_rp_r_c_1c))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_satz253 x (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz250 (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)))). Time Defined. (* constant 5348 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov y (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n))) l_e_st_eq_landau_n_rt_rp_r_1rl))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz270 (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_ov y (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t4 x y n)))). Time Defined. (* constant 5349 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_fx : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_ov (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t3 x y n)))). Time Defined. (* constant 5350 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_fy : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_ov (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t3 x y n)))). Time Defined. (* constant 5351 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) l_e_st_eq_landau_n_rt_rp_r_1rl))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_ismoreis1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov y (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov x (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ov y (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n) (l_e_st_eq_landau_n_rt_rp_r_c_satz269 x (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_satz269 y (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t5 x y n)))). Time Defined. (* constant 5352 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_prl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))))). Time Defined. (* constant 5353 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_prr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))))). Time Defined. (* constant 5354 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_7271_prl x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_prr x y n)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_orapp (l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_7271_prl x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_prr x y n)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t6 x y n) (fun (t:l_e_st_eq_landau_n_rt_rp_r_more (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) l_e_st_eq_landau_n_rt_rp_r_1rl) => l_e_st_eq_landau_n_rt_rp_r_moreisi1 (l_e_st_eq_landau_n_rt_rp_r_c_7271_prl x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_prr x y n) (l_e_st_eq_landau_n_rt_rp_r_satz203a (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) t (l_e_st_eq_landau_n_rt_rp_r_c_satz264a (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) n))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) l_e_st_eq_landau_n_rt_rp_r_1rl) => l_e_st_eq_landau_n_rt_rp_r_moreisi2 (l_e_st_eq_landau_n_rt_rp_r_c_7271_prl x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_prr x y n) (l_e_st_eq_landau_n_rt_rp_r_ists1 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n)) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) t))))). Time Defined. (* constant 5355 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_7271_prl x y n) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_7271_prl x y n) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)))) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_disttp1 (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_7271_fx x y n) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))) (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_7271_fy x y n) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))) (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_satz204e (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t3 x y n)) (l_e_st_eq_landau_n_rt_rp_r_satz204e (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t3 x y n)))))). Time Defined. (* constant 5356 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_7271_prr x y n) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_satz195b (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))))). Time Defined. (* constant 5357 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_ismoreis12 (l_e_st_eq_landau_n_rt_rp_r_c_7271_prl x y n) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_prr x y n) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t8 x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t9 x y n) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t7 x y n)))). Time Defined. (* constant 5358 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7271_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_7271_t2 x y t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pl x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_7271_t10 x y t))). Time Defined. (* constant 5359 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz271 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_satz168a (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)) (l_e_st_eq_landau_n_rt_rp_r_c_7271_t11 x y))). Time Defined. (* constant 5360 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz271a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_7271_t11 x y)). Time Defined. (* constant 5361 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7272_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_iscere (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_satz214 x)) (l_e_st_eq_landau_n_rt_rp_r_c_reis (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). Time Defined. (* constant 5362 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7272_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_7272_t1 x) (l_e_st_eq_landau_n_rt_rp_r_c_7272_t1 x)) (l_e_st_eq_landau_n_rt_rp_r_satz198 (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x))). Time Defined. (* constant 5363 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7272_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_isceim (l_e_st_eq_landau_n_rt_rp_r_c_m0 x) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_c_satz214 x)) (l_e_st_eq_landau_n_rt_rp_r_c_imis (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)))). Time Defined. (* constant 5364 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7272_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_7272_t3 x) (l_e_st_eq_landau_n_rt_rp_r_c_7272_t3 x)) (l_e_st_eq_landau_n_rt_rp_r_satz198 (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))). Time Defined. (* constant 5365 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7272_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod2 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_m0 x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_m0 x))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_7272_t2 x) (l_e_st_eq_landau_n_rt_rp_r_c_7272_t4 x)). Time Defined. (* constant 5366 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz272 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_issqrt (l_e_st_eq_landau_n_rt_rp_r_c_mod2 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_mod2 x) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 x) (l_e_st_eq_landau_n_rt_rp_r_c_7272_t5 x)). Time Defined. (* constant 5367 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz272a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_m0 x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_m0 x)) (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_satz272 x)). Time Defined. (* constant 5368 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7273_sum : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)))). Time Defined. (* constant 5369 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7273_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_islessis1 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_mn x y))) (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_c_ismod (l_e_st_eq_landau_n_rt_rp_r_c_pl y (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) x (l_e_st_eq_landau_n_rt_rp_r_c_satz212h x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz271 y (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)))). Time Defined. (* constant 5370 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7273_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y)) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) (l_e_st_eq_landau_n_rt_rp_r_c_7273_t1 x y) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y)) => l_e_st_eq_landau_n_rt_rp_r_satz188f (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y)) => l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_c_mod y) t))). Time Defined. (* constant 5371 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7273_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)))) (l_e_st_eq_landau_n_rt_rp_r_mnpl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)))). Time Defined. (* constant 5372 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7273_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_satz168b (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_islessis2 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_7273_sum x y) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_7273_t3 x y) (l_e_st_eq_landau_n_rt_rp_r_c_7273_t2 x y)))). Time Defined. (* constant 5373 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7273_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ismoreis12 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn y x)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))) (l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn y x)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y))) (l_e_st_eq_landau_n_rt_rp_r_c_ismod (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_c_mn y x) (l_e_st_eq_landau_n_rt_rp_r_c_satz219 x y)) (l_e_st_eq_landau_n_rt_rp_r_c_satz272 (l_e_st_eq_landau_n_rt_rp_r_c_mn x y))) (l_e_st_eq_landau_n_rt_rp_r_satz181a (l_e_st_eq_landau_n_rt_rp_r_c_mod y) (l_e_st_eq_landau_n_rt_rp_r_c_mod x)) (l_e_st_eq_landau_n_rt_rp_r_c_7273_t4 y x))). Time Defined. (* constant 5374 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7273_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_or (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_neg r) (l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_m0 r)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) r) (l_or_th6 (l_e_st_eq_landau_n_rt_rp_r_neg r)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_absn r t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg r)) => l_e_st_eq_landau_n_rt_rp_r_absnn r t)). Time Defined. (* constant 5375 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_7273_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s), (forall (n:l_e_st_eq_landau_n_rt_rp_r_moreis r (l_e_st_eq_landau_n_rt_rp_r_m0 s)), l_e_st_eq_landau_n_rt_rp_r_moreis r (l_e_st_eq_landau_n_rt_rp_r_abs s))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_moreis r s) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_moreis r (l_e_st_eq_landau_n_rt_rp_r_m0 s)) => l_orapp (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs s) s) (l_e_st_eq_landau_n_rt_rp_r_moreis r (l_e_st_eq_landau_n_rt_rp_r_abs s)) (l_e_st_eq_landau_n_rt_rp_r_c_7273_t6 s) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_r_m0 s)) => l_e_st_eq_landau_n_rt_rp_r_ismoreis2 (l_e_st_eq_landau_n_rt_rp_r_m0 s) (l_e_st_eq_landau_n_rt_rp_r_abs s) r (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs s) (l_e_st_eq_landau_n_rt_rp_r_m0 s) t) n) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs s) s) => l_e_st_eq_landau_n_rt_rp_r_ismoreis2 s (l_e_st_eq_landau_n_rt_rp_r_abs s) r (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs s) s t) m))))). Time Defined. (* constant 5376 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz273 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_moreis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_7273_t7 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_mn x y)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_mod x) (l_e_st_eq_landau_n_rt_rp_r_c_mod y)) (l_e_st_eq_landau_n_rt_rp_r_c_7273_t4 x y) (l_e_st_eq_landau_n_rt_rp_r_c_7273_t5 x y))). Time Defined. (* constant 5377 *) Definition l_e_st_eq_landau_n_8274_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_some (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to y) (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to y)) => l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to y) f))). Time Defined. (* constant 5378 *) Definition l_e_st_eq_landau_n_8274_prop2 : (forall (x:l_e_st_eq_landau_n_nat), Prop). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_all (fun (y:l_e_st_eq_landau_n_nat) => l_imp (l_e_st_eq_landau_n_less x y) (l_not (l_e_st_eq_landau_n_8274_prop1 x y)))). Time Defined. (* constant 5379 *) Definition l_e_st_eq_landau_n_8274_1y : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), l_e_st_eq_landau_n_1to y))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => l_e_st_eq_landau_n_1out y))). Time Defined. (* constant 5380 *) Definition l_e_st_eq_landau_n_8274_yy : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), l_e_st_eq_landau_n_1to y))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => l_e_st_eq_landau_n_xout y))). Time Defined. (* constant 5381 *) Definition l_e_st_eq_landau_n_8274_t1 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_yy y l f)), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 y)))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_yy y l f)) => l_e_st_eq_landau_n_isoutne y l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a y) y (l_e_st_eq_landau_n_lessisi3 y) i)))). Time Defined. (* constant 5382 *) Definition l_e_st_eq_landau_n_8274_t2 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), l_e_st_eq_landau_n_nis l_e_st_eq_landau_n_1 y))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => l_ec3e31 (l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 y) (l_e_st_eq_landau_n_more l_e_st_eq_landau_n_1 y) (l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) (l_e_st_eq_landau_n_satz10b l_e_st_eq_landau_n_1 y) l))). Time Defined. (* constant 5383 *) Definition l_e_st_eq_landau_n_8274_t3 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), l_not (l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_yy y l f))))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_yy y l f)) (l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 y) (l_e_st_eq_landau_n_8274_t2 y l f) (fun (t:l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_yy y l f)) => l_e_st_eq_landau_n_8274_t1 y l f t)))). Time Defined. (* constant 5384 *) Definition l_e_st_eq_landau_n_8274_t4 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to y) (f u) (f l_e_st_eq_landau_n_1o))))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f u l_e_st_eq_landau_n_1o (l_e_st_eq_landau_n_singlet_th1 u))))). Time Defined. (* constant 5385 *) Definition l_e_st_eq_landau_n_8274_t5 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f)), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f u) (l_e_st_eq_landau_n_8274_yy y l f))))))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f)) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_notis_th2 (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_yy y l f) (f u) (l_e_st_eq_landau_n_8274_t3 y l f) (l_e_tris (l_e_st_eq_landau_n_1to y) (f u) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_t4 y l f u) i)))))). Time Defined. (* constant 5386 *) Definition l_e_st_eq_landau_n_8274_t6 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f)), l_not (l_e_image (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f (l_e_st_eq_landau_n_8274_yy y l f)))))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f)) => l_some_th5 (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_yy y l f) (f u)) (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_symnotis (l_e_st_eq_landau_n_1to y) (f u) (l_e_st_eq_landau_n_8274_yy y l f) (l_e_st_eq_landau_n_8274_t5 y l f i u)))))). Time Defined. (* constant 5387 *) Definition l_e_st_eq_landau_n_8274_t7 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f)), l_not (l_e_surjective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f))))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f)) => l_all_th1 (l_e_st_eq_landau_n_1to y) (fun (u:l_e_st_eq_landau_n_1to y) => l_e_image (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f u) (l_e_st_eq_landau_n_8274_yy y l f) (l_e_st_eq_landau_n_8274_t6 y l f i))))). Time Defined. (* constant 5388 *) Definition l_e_st_eq_landau_n_8274_t8 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f))), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f u) (l_e_st_eq_landau_n_8274_1y y l f))))))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f))) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_notis_th2 (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f) (f u) n (l_e_st_eq_landau_n_8274_t4 y l f u)))))). Time Defined. (* constant 5389 *) Definition l_e_st_eq_landau_n_8274_t9 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f))), l_not (l_e_image (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f (l_e_st_eq_landau_n_8274_1y y l f)))))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f))) => l_some_th5 (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_1y y l f) (f u)) (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_symnotis (l_e_st_eq_landau_n_1to y) (f u) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_t8 y l f n u)))))). Time Defined. (* constant 5390 *) Definition l_e_st_eq_landau_n_8274_t10 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f))), l_not (l_e_surjective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f))))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f))) => l_all_th1 (l_e_st_eq_landau_n_1to y) (fun (u:l_e_st_eq_landau_n_1to y) => l_e_image (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f u) (l_e_st_eq_landau_n_8274_1y y l f) (l_e_st_eq_landau_n_8274_t9 y l f n))))). Time Defined. (* constant 5391 *) Definition l_e_st_eq_landau_n_8274_t11 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), l_not (l_e_surjective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f)))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f)) (l_not (l_e_surjective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f)) (fun (t:l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f)) => l_e_st_eq_landau_n_8274_t7 y l f t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f l_e_st_eq_landau_n_1o) (l_e_st_eq_landau_n_8274_1y y l f))) => l_e_st_eq_landau_n_8274_t10 y l f t)))). Time Defined. (* constant 5392 *) Definition l_e_st_eq_landau_n_8274_t12 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)), l_not (l_e_bijective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f)))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => l_and_th2 (l_e_injective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f) (l_e_surjective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f) (l_e_st_eq_landau_n_8274_t11 y l f)))). Time Defined. (* constant 5393 *) Definition l_e_st_eq_landau_n_8274_t13 : (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y), l_not (l_e_st_eq_landau_n_8274_prop1 l_e_st_eq_landau_n_1 y))). exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => l_some_th5 (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y) (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => l_e_bijective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to y) f) (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to y)) => l_e_st_eq_landau_n_8274_t12 y l f))). Time Defined. (* constant 5394 *) Definition l_e_st_eq_landau_n_8274_t14 : l_e_st_eq_landau_n_8274_prop2 l_e_st_eq_landau_n_1. exact (fun (y:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y) => l_e_st_eq_landau_n_8274_t13 y t)). Time Defined. (* constant 5395 *) Definition l_e_st_eq_landau_n_8274_xs : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_suc x). Time Defined. (* constant 5396 *) Definition l_e_st_eq_landau_n_8274_xxs : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_suc x))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_8274_xs x))))). Time Defined. (* constant 5397 *) Definition l_e_st_eq_landau_n_8274_yy1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), l_e_st_eq_landau_n_1to y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => l_e_st_eq_landau_n_xout y)))). Time Defined. (* constant 5398 *) Definition l_e_st_eq_landau_n_8274_t15 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), l_e_st_eq_landau_n_less l_e_st_eq_landau_n_1 y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => l_e_st_eq_landau_n_trless l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_suc x) y (l_e_st_eq_landau_n_satz24c x) l)))). Time Defined. (* constant 5399 *) Definition l_e_st_eq_landau_n_8274_ym1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), l_e_st_eq_landau_n_nat)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => l_e_st_eq_landau_n_mn y l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_8274_t15 x p y l))))). Time Defined. (* constant 5400 *) Definition l_e_st_eq_landau_n_8274_t16 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => l_e_st_eq_landau_n_isless12 (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) y (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz4e x) (l_e_st_eq_landau_n_mn_th1c y l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_8274_t15 x p y l)) l)))). Time Defined. (* constant 5401 *) Definition l_e_st_eq_landau_n_8274_t17 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), l_e_st_eq_landau_n_less x (l_e_st_eq_landau_n_8274_ym1 x p y l))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => l_e_st_eq_landau_n_satz20c x (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_8274_t16 x p y l))))). Time Defined. (* constant 5402 *) Definition l_e_st_eq_landau_n_8274_t18 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_ym1 x p y l) y)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => l_e_st_eq_landau_n_isless2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1) y (l_e_st_eq_landau_n_8274_ym1 x p y l) (l_e_st_eq_landau_n_mn_th1d y l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_8274_t15 x p y l)) (l_e_st_eq_landau_n_satz18a (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1))))). Time Defined. (* constant 5403 *) Definition l_e_st_eq_landau_n_8274_t19 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), l_e_injective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => l_ande1 (l_e_injective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) (l_e_surjective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) b)))))). Time Defined. (* constant 5404 *) Definition l_e_st_eq_landau_n_8274_t20 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), l_e_surjective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => l_ande2 (l_e_injective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) (l_e_surjective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) b)))))). Time Defined. (* constant 5405 *) Definition l_e_st_eq_landau_n_8274_u1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_inn x u)))))))). Time Defined. (* constant 5406 *) Definition l_e_st_eq_landau_n_8274_t21 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_satz16a (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) x (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_1top x u) (l_e_st_eq_landau_n_satz18c x))))))))). Time Defined. (* constant 5407 *) Definition l_e_st_eq_landau_n_8274_t22 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_ec3e31 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_t21 x p y l f b i u))))))))). Time Defined. (* constant 5408 *) Definition l_e_st_eq_landau_n_8274_t23 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_lessisi1 (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_t21 x p y l f b i u))))))))). Time Defined. (* constant 5409 *) Definition l_e_st_eq_landau_n_8274_u2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_t23 x p y l f b i u))))))))). Time Defined. (* constant 5410 *) Definition l_e_st_eq_landau_n_8274_t24 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_u2 x p y l f b i u) (l_e_st_eq_landau_n_8274_xxs x p y l)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_u2 x p y l f b i u) (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_t22 x p y l f b i u) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_u2 x p y l f b i u) (l_e_st_eq_landau_n_8274_xxs x p y l)) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_t23 x p y l f b i u) (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_8274_xs x)) t))))))))). Time Defined. (* constant 5411 *) Definition l_e_st_eq_landau_n_8274_t25 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), (forall (j:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (f (l_e_st_eq_landau_n_8274_xxs x p y l))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => l_e_tris2 (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l) j i))))))))). Time Defined. (* constant 5412 *) Definition l_e_st_eq_landau_n_8274_t26 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), (forall (j:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_u2 x p y l f b i u) (l_e_st_eq_landau_n_8274_xxs x p y l)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => l_e_isfe (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f (l_e_st_eq_landau_n_8274_t19 x p y l f b) (l_e_st_eq_landau_n_8274_u2 x p y l f b i u) (l_e_st_eq_landau_n_8274_xxs x p y l) (l_e_st_eq_landau_n_8274_t25 x p y l f b i u j)))))))))). Time Defined. (* constant 5413 *) Definition l_e_st_eq_landau_n_8274_t27 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_u2 x p y l f b i u) (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_t24 x p y l f b i u) (fun (t:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => l_e_st_eq_landau_n_8274_t26 x p y l f b i u t))))))))). Time Defined. (* constant 5414 *) Definition l_e_st_eq_landau_n_8274_w1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_inn y (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)))))))))). Time Defined. (* constant 5415 *) Definition l_e_st_eq_landau_n_8274_t28 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) y), l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) y) => l_e_tris (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_outn y (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_1top y (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)))) (l_e_st_eq_landau_n_8274_yy1 x p y l) (l_e_st_eq_landau_n_isoutinn y (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u))) (l_e_st_eq_landau_n_isoutni y (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_1top y (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u))) y (l_e_st_eq_landau_n_lessisi3 y) j)))))))))). Time Defined. (* constant 5416 *) Definition l_e_st_eq_landau_n_8274_t29 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_imp_th3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) y) (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) (l_e_st_eq_landau_n_8274_t27 x p y l f b i u) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) y) => l_e_st_eq_landau_n_8274_t28 x p y l f b i u t))))))))). Time Defined. (* constant 5417 *) Definition l_e_st_eq_landau_n_8274_t30 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_ore1 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) y) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) y) (l_e_st_eq_landau_n_1top y (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u))) (l_e_st_eq_landau_n_8274_t29 x p y l f b i u))))))))). Time Defined. (* constant 5418 *) Definition l_e_st_eq_landau_n_8274_t31 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_islessis2 y (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_mn_th1c y l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_8274_t15 x p y l)) (l_e_st_eq_landau_n_satz25b y (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_t30 x p y l f b i u)))))))))). Time Defined. (* constant 5419 *) Definition l_e_st_eq_landau_n_8274_t32 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_ym1 x p y l))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_or_th9 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_t31 x p y l f b i u) (fun (t:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_satz20c (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1 t) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_satz20b (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_ym1 x p y l) l_e_st_eq_landau_n_1 t))))))))). Time Defined. (* constant 5420 *) Definition l_e_st_eq_landau_n_8274_w2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_8274_ym1 x p y l) (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_t32 x p y l f b i u))))))))). Time Defined. (* constant 5421 *) Definition l_e_st_eq_landau_n_8274_f1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_8274_w2 x p y l f b i t)))))))). Time Defined. (* constant 5422 *) Definition l_e_st_eq_landau_n_8274_t33 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_w1 x p y l f b i v))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_8274_ym1 x p y l) (l_e_st_eq_landau_n_8274_w1 x p y l f b i u) (l_e_st_eq_landau_n_8274_t32 x p y l f b i u) (l_e_st_eq_landau_n_8274_w1 x p y l f b i v) (l_e_st_eq_landau_n_8274_t32 x p y l f b i v) j)))))))))). Time Defined. (* constant 5423 *) Definition l_e_st_eq_landau_n_8274_t34 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)), l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i v)))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)) => l_e_st_eq_landau_n_isinne y (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i u)) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i v)) (l_e_st_eq_landau_n_8274_t33 x p y l f b i u v j))))))))))). Time Defined. (* constant 5424 *) Definition l_e_st_eq_landau_n_8274_t35 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_u2 x p y l f b i u) (l_e_st_eq_landau_n_8274_u2 x p y l f b i v))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)) => l_e_st_eq_landau_n_8274_t19 x p y l f b (l_e_st_eq_landau_n_8274_u2 x p y l f b i u) (l_e_st_eq_landau_n_8274_u2 x p y l f b i v) (l_e_st_eq_landau_n_8274_t34 x p y l f b i u v j))))))))))). Time Defined. (* constant 5425 *) Definition l_e_st_eq_landau_n_8274_t36 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_u1 x p y l f b i v))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_u1 x p y l f b i u) (l_e_st_eq_landau_n_8274_t23 x p y l f b i u) (l_e_st_eq_landau_n_8274_u1 x p y l f b i v) (l_e_st_eq_landau_n_8274_t23 x p y l f b i v) (l_e_st_eq_landau_n_8274_t35 x p y l f b i u v j))))))))))). Time Defined. (* constant 5426 *) Definition l_e_st_eq_landau_n_8274_t37 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)), l_e_is (l_e_st_eq_landau_n_1to x) u v)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)) => l_e_st_eq_landau_n_isinne x u v (l_e_st_eq_landau_n_8274_t36 x p y l f b i u v j))))))))))). Time Defined. (* constant 5427 *) Definition l_e_st_eq_landau_n_8274_v1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_nat)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_inn (l_e_st_eq_landau_n_8274_ym1 x p y l) v)))))))). Time Defined. (* constant 5428 *) Definition l_e_st_eq_landau_n_8274_t38 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_satz16a (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) (l_e_st_eq_landau_n_8274_ym1 x p y l) y (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_8274_ym1 x p y l) v) (l_e_st_eq_landau_n_8274_t18 x p y l))))))))). Time Defined. (* constant 5429 *) Definition l_e_st_eq_landau_n_8274_t39 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_ec3e31 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y) (l_e_st_eq_landau_n_8274_t38 x p y l f b i v))))))))). Time Defined. (* constant 5430 *) Definition l_e_st_eq_landau_n_8274_t40 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_lessisi1 (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y (l_e_st_eq_landau_n_8274_t38 x p y l f b i v))))))))). Time Defined. (* constant 5431 *) Definition l_e_st_eq_landau_n_8274_v2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_1to y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_outn y (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) (l_e_st_eq_landau_n_8274_t40 x p y l f b i v))))))))). Time Defined. (* constant 5432 *) Definition l_e_st_eq_landau_n_8274_w3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_invf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f b (l_e_st_eq_landau_n_8274_v2 x p y l f b i v))))))))). Time Defined. (* constant 5433 *) Definition l_e_st_eq_landau_n_8274_t41 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_v2 x p y l f b i v) (f (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_thinvf2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f b (l_e_st_eq_landau_n_8274_v2 x p y l f b i v))))))))). Time Defined. (* constant 5434 *) Definition l_e_st_eq_landau_n_8274_t42 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)), l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (f (l_e_st_eq_landau_n_8274_xxs x p y l))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l) j))))))))). Time Defined. (* constant 5435 *) Definition l_e_st_eq_landau_n_8274_t43 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)), l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_v2 x p y l f b i v) (l_e_st_eq_landau_n_8274_yy1 x p y l)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)) => l_e_tr3is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_v2 x p y l f b i v) (f (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l) (l_e_st_eq_landau_n_8274_t41 x p y l f b i v) (l_e_st_eq_landau_n_8274_t42 x p y l f b i v j) i))))))))). Time Defined. (* constant 5436 *) Definition l_e_st_eq_landau_n_8274_t44 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)) => l_e_st_eq_landau_n_isoutne y (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) (l_e_st_eq_landau_n_8274_t40 x p y l f b i v) y (l_e_st_eq_landau_n_lessisi3 y) (l_e_st_eq_landau_n_8274_t43 x p y l f b i v j)))))))))). Time Defined. (* constant 5437 *) Definition l_e_st_eq_landau_n_8274_t45 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) y) (l_e_st_eq_landau_n_8274_t39 x p y l f b i v) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)) => l_e_st_eq_landau_n_8274_t44 x p y l f b i v t))))))))). Time Defined. (* constant 5438 *) Definition l_e_st_eq_landau_n_8274_w4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_nat)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_inn (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v))))))))). Time Defined. (* constant 5439 *) Definition l_e_st_eq_landau_n_8274_t46 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_xs x)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_xs x)) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v))) (l_e_st_eq_landau_n_8274_xxs x p y l) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_8274_xs x)) j)))))))))). Time Defined. (* constant 5440 *) Definition l_e_st_eq_landau_n_8274_t47 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_xs x))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_imp_th3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_xs x)) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_t45 x p y l f b i v) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_xs x)) => l_e_st_eq_landau_n_8274_t46 x p y l f b i v t))))))))). Time Defined. (* constant 5441 *) Definition l_e_st_eq_landau_n_8274_t48 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_xs x))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_ore1 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (l_e_st_eq_landau_n_8274_t47 x p y l f b i v))))))))). Time Defined. (* constant 5442 *) Definition l_e_st_eq_landau_n_8274_t49 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) x)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_satz26a x (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_t48 x p y l f b i v))))))))). Time Defined. (* constant 5443 *) Definition l_e_st_eq_landau_n_8274_w5 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_1to x)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_t49 x p y l f b i v))))))))). Time Defined. (* constant 5444 *) Definition l_e_st_eq_landau_n_8274_t50 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_u1 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_8274_t49 x p y l f b i v))))))))). Time Defined. (* constant 5445 *) Definition l_e_st_eq_landau_n_8274_t51 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_u2 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v))) (l_e_st_eq_landau_n_8274_u2 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w4 x p y l f b i v) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_8274_xs x) (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (l_e_st_eq_landau_n_8274_u1 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)) (l_e_st_eq_landau_n_8274_t23 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)) (l_e_st_eq_landau_n_8274_t50 x p y l f b i v)))))))))). Time Defined. (* constant 5446 *) Definition l_e_st_eq_landau_n_8274_t52 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f (l_e_st_eq_landau_n_8274_w3 x p y l f b i v) (l_e_st_eq_landau_n_8274_u2 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)) (l_e_st_eq_landau_n_8274_t51 x p y l f b i v))))))))). Time Defined. (* constant 5447 *) Definition l_e_st_eq_landau_n_8274_t53 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_v2 x p y l f b i v) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v))))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_tris (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_v2 x p y l f b i v) (f (l_e_st_eq_landau_n_8274_w3 x p y l f b i v)) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v))) (l_e_st_eq_landau_n_8274_t41 x p y l f b i v) (l_e_st_eq_landau_n_8274_t52 x p y l f b i v))))))))). Time Defined. (* constant 5448 *) Definition l_e_st_eq_landau_n_8274_t54 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) (l_e_st_eq_landau_n_8274_w1 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) (l_e_st_eq_landau_n_inn y (l_e_st_eq_landau_n_8274_v2 x p y l f b i v)) (l_e_st_eq_landau_n_8274_w1 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)) (l_e_st_eq_landau_n_isinoutn y (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) (l_e_st_eq_landau_n_8274_t40 x p y l f b i v)) (l_e_st_eq_landau_n_isinni y (l_e_st_eq_landau_n_8274_v2 x p y l f b i v) (f (l_e_st_eq_landau_n_8274_u2 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v))) (l_e_st_eq_landau_n_8274_t53 x p y l f b i v)))))))))). Time Defined. (* constant 5449 *) Definition l_e_st_eq_landau_n_8274_t55 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) v (l_e_st_eq_landau_n_8274_f1 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) v (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_8274_ym1 x p y l) (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_8274_ym1 x p y l) v)) (l_e_st_eq_landau_n_8274_w2 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_8274_ym1 x p y l) v) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_8274_ym1 x p y l) (l_e_st_eq_landau_n_8274_v1 x p y l f b i v) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_8274_ym1 x p y l) v) (l_e_st_eq_landau_n_8274_w1 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)) (l_e_st_eq_landau_n_8274_t32 x p y l f b i (l_e_st_eq_landau_n_8274_w5 x p y l f b i v)) (l_e_st_eq_landau_n_8274_t54 x p y l f b i v)))))))))). Time Defined. (* constant 5450 *) Definition l_e_st_eq_landau_n_8274_t56 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), (forall (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)), l_e_image (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i) v)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => (fun (v:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_somei (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) v (l_e_st_eq_landau_n_8274_f1 x p y l f b i t)) (l_e_st_eq_landau_n_8274_w5 x p y l f b i v) (l_e_st_eq_landau_n_8274_t55 x p y l f b i v))))))))). Time Defined. (* constant 5451 *) Definition l_e_st_eq_landau_n_8274_t57 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => l_andi (l_e_injective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i)) (l_e_surjective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i)) (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (l_e_st_eq_landau_n_8274_f1 x p y l f b i u) (l_e_st_eq_landau_n_8274_f1 x p y l f b i v)) => l_e_st_eq_landau_n_8274_t37 x p y l f b i u v t))) (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) => l_e_st_eq_landau_n_8274_t56 x p y l f b i u)))))))). Time Defined. (* constant 5452 *) Definition l_e_st_eq_landau_n_8274_t58 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), l_e_st_eq_landau_n_8274_prop1 x (l_e_st_eq_landau_n_8274_ym1 x p y l)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => l_somei (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l))) => l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_ym1 x p y l)) g) (l_e_st_eq_landau_n_8274_f1 x p y l f b i) (l_e_st_eq_landau_n_8274_t57 x p y l f b i)))))))). Time Defined. (* constant 5453 *) Definition l_e_st_eq_landau_n_8274_t59 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)), l_con))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => p (l_e_st_eq_landau_n_8274_ym1 x p y l) (l_e_st_eq_landau_n_8274_t17 x p y l) (l_e_st_eq_landau_n_8274_t58 x p y l f b i)))))))). Time Defined. (* constant 5454 *) Definition l_e_st_eq_landau_n_8274_m0 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))) => l_e_invf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f b (l_e_st_eq_landau_n_8274_yy1 x p y l)))))))). Time Defined. (* constant 5455 *) Definition l_e_st_eq_landau_n_8274_t60 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))), l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_yy1 x p y l) (f (l_e_st_eq_landau_n_8274_m0 x p y l f b n))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))) => l_e_thinvf2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f b (l_e_st_eq_landau_n_8274_yy1 x p y l)))))))). Time Defined. (* constant 5456 *) Definition l_e_st_eq_landau_n_8274_f2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))) => l_e_changef (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f (l_e_st_eq_landau_n_8274_m0 x p y l f b n) (l_e_st_eq_landau_n_8274_xxs x p y l)))))))). Time Defined. (* constant 5457 *) Definition l_e_st_eq_landau_n_8274_t61 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))), l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_f2 x p y l f b n (l_e_st_eq_landau_n_8274_xxs x p y l)) (f (l_e_st_eq_landau_n_8274_m0 x p y l f b n))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))) => l_e_changef2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f (l_e_st_eq_landau_n_8274_m0 x p y l f b n) (l_e_st_eq_landau_n_8274_xxs x p y l) (l_e_st_eq_landau_n_8274_xxs x p y l) (l_e_refis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_8274_xxs x p y l))))))))). Time Defined. (* constant 5458 *) Definition l_e_st_eq_landau_n_8274_t62 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))), l_e_is (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_f2 x p y l f b n (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))) => l_e_tris2 (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_f2 x p y l f b n (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l) (f (l_e_st_eq_landau_n_8274_m0 x p y l f b n)) (l_e_st_eq_landau_n_8274_t61 x p y l f b n) (l_e_st_eq_landau_n_8274_t60 x p y l f b n)))))))). Time Defined. (* constant 5459 *) Definition l_e_st_eq_landau_n_8274_t63 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))), l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) (l_e_st_eq_landau_n_8274_f2 x p y l f b n)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))) => l_e_wissel_th6 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f (l_e_st_eq_landau_n_8274_m0 x p y l f b n) (l_e_st_eq_landau_n_8274_xxs x p y l) b))))))). Time Defined. (* constant 5460 *) Definition l_e_st_eq_landau_n_8274_t64 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))), l_con))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))) => l_e_st_eq_landau_n_8274_t59 x p y l (l_e_st_eq_landau_n_8274_f2 x p y l f b n) (l_e_st_eq_landau_n_8274_t63 x p y l f b n) (l_e_st_eq_landau_n_8274_t62 x p y l f b n)))))))). Time Defined. (* constant 5461 *) Definition l_e_st_eq_landau_n_8274_t65 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f), l_con)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) l_con (fun (t:l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l)) => l_e_st_eq_landau_n_8274_t59 x p y l f b t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to y) (f (l_e_st_eq_landau_n_8274_xxs x p y l)) (l_e_st_eq_landau_n_8274_yy1 x p y l))) => l_e_st_eq_landau_n_8274_t64 x p y l f b t))))))). Time Defined. (* constant 5462 *) Definition l_e_st_eq_landau_n_8274_t65a : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y), l_not (l_e_st_eq_landau_n_8274_prop1 (l_e_st_eq_landau_n_8274_xs x) y))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_suc x) y) => l_some_th5 (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y) (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)), l_e_st_eq_landau_n_1to y)) => (fun (t:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_8274_xs x)) (l_e_st_eq_landau_n_1to y) f) => l_e_st_eq_landau_n_8274_t65 x p y l f t)))))). Time Defined. (* constant 5463 *) Definition l_e_st_eq_landau_n_8274_t66 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_8274_prop2 x), l_e_st_eq_landau_n_8274_prop2 (l_e_st_eq_landau_n_8274_xs x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_8274_prop2 x) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_8274_xs x) y) => l_e_st_eq_landau_n_8274_t65a x p y t)))). Time Defined. (* constant 5464 *) Definition l_e_st_eq_landau_n_8274_t67 : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_8274_prop2 x). exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_8274_prop2 t) l_e_st_eq_landau_n_8274_t14 (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_8274_prop2 t) => l_e_st_eq_landau_n_8274_t66 t u)) x). Time Defined. (* constant 5465 *) Definition l_e_st_eq_landau_n_satz274 : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), l_not (l_some (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to y) (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to y)) => l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to y) f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => l_e_st_eq_landau_n_8274_t67 x y l))). Time Defined. (* constant 5466 *) Definition l_e_st_eq_landau_n_satz274a : (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_less x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to y)), l_not (l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to y) f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_less x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to y)) => l_some_th4 (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to y) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to y)) => l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to y) g) (l_e_st_eq_landau_n_satz274 x y l) f)))). Time Defined. (* constant 5467 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_inn : (forall (x:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_inn x u)). Time Defined. (* constant 5468 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_1to x), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))), l_not (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))) => l_imp_th3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x) (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x)) o (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x) => l_e_tris (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n)) (l_e_st_eq_landau_n_xout x) (l_e_st_eq_landau_n_isoutinn x n) (l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n) x (l_e_st_eq_landau_n_lessisi3 x) t)))))). Time Defined. (* constant 5469 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_1to x), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))) => l_ore1 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x) (l_e_st_eq_landau_n_1top x n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t1 q x n o))))). Time Defined. (* constant 5470 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma275 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_1to x), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x n)) x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))) => l_e_st_eq_landau_n_satz25c x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t2 q x n o))))). Time Defined. (* constant 5471 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_recprop : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_and (l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_1out x)) (f (l_e_st_eq_landau_n_1out x))) (forall (t:l_e_st_eq_landau_n_1to x), (forall (u:l_not (l_e_is (l_e_st_eq_landau_n_1to x) t (l_e_st_eq_landau_n_xout x))), l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x t)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q x t u))) (q (g t) (f (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x t)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q x t u)))))))))). Time Defined. (* constant 5472 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_1o : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_1to x)). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_1out x)). Time Defined. (* constant 5473 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_xo : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_1to x)). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_xout x)). Time Defined. (* constant 5474 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t11 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_less u x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_st_eq_landau_n_satz16b u (l_e_st_eq_landau_n_suc u) x (l_e_st_eq_landau_n_satz18c u) l)))). Time Defined. (* constant 5475 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t12 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_lessis u x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_st_eq_landau_n_lessisi1 u x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t11 q x u l))))). Time Defined. (* constant 5476 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_ux : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_1to x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_st_eq_landau_n_outn x u (l_e_st_eq_landau_n_rt_rp_r_c_8275_t12 q x u l))))). Time Defined. (* constant 5477 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t13 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_nis u x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_ec3e31 (l_e_st_eq_landau_n_is u x) (l_e_st_eq_landau_n_more u x) (l_e_st_eq_landau_n_less u x) (l_e_st_eq_landau_n_satz10b u x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t11 q x u l))))). Time Defined. (* constant 5478 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_not (l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) (l_e_st_eq_landau_n_is u x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t13 q x u l) (fun (t:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) => l_e_st_eq_landau_n_isoutne x u (l_e_st_eq_landau_n_rt_rp_r_c_8275_t12 q x u l) x (l_e_st_eq_landau_n_lessisi3 x) t))))). Time Defined. (* constant 5479 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t15 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc u) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_isf l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_suc u (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (l_e_st_eq_landau_n_isinoutn x u (l_e_st_eq_landau_n_rt_rp_r_c_8275_t12 q x u l)))))). Time Defined. (* constant 5480 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t16 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l) (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l))) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_suc u) l (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l))) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t15 q x u l))))). Time Defined. (* constant 5481 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_ns : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (n:l_e_st_eq_landau_n_1to x), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x))), l_e_st_eq_landau_n_1to x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x))) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q x n o))))). Time Defined. (* constant 5482 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)))))). Time Defined. (* constant 5483 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_prop2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (forall (t:l_e_st_eq_landau_n_1to x), (forall (u:l_not (l_e_is (l_e_st_eq_landau_n_1to x) t (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x))), l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x t u)) (q (g t) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x t u))))))))). Time Defined. (* constant 5484 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q x f g))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q x f g) (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop2 q x f g) pg))))). Time Defined. (* constant 5485 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (n:l_e_st_eq_landau_n_1to x), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x))), l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x n o)) (q (g n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x n o)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x))) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q x f g) (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop2 q x f g) pg n o))))))). Time Defined. (* constant 5486 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t17 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l)) (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_isf (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx g (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t16 q x u l)))))))). Time Defined. (* constant 5487 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t18 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l)) (q (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (g (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l)) (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l))) (q (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t17 q x f g pg u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t4 q x f g pg (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l))))))))). Time Defined. (* constant 5488 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (u:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis u x), Prop))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis u x) => l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_outn x u l)) (h (l_e_st_eq_landau_n_outn x u l))))))))). Time Defined. (* constant 5489 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (u:l_e_st_eq_landau_n_nat), Prop)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (u:l_e_st_eq_landau_n_nat) => l_and (l_e_st_eq_landau_n_lessis u x) (forall (t:l_e_st_eq_landau_n_lessis u x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h u t))))))). Time Defined. (* constant 5490 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (u:l_e_st_eq_landau_n_nat), Prop)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (u:l_e_st_eq_landau_n_nat) => l_or (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h u) (l_e_st_eq_landau_n_more u x))))))). Time Defined. (* constant 5491 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (l:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_outn x l_e_st_eq_landau_n_1 l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (l:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x) => l_e_st_eq_landau_n_isoutni x l_e_st_eq_landau_n_1 l l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a x) (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1))))))))). Time Defined. (* constant 5492 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t6 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (l:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x), l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_outn x l_e_st_eq_landau_n_1 l)) (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (l:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x) => l_e_isf (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx g (l_e_st_eq_landau_n_outn x l_e_st_eq_landau_n_1 l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t5 q x f g h pg ph l))))))))). Time Defined. (* constant 5493 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t7 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (l:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x), l_e_st_eq_landau_n_rt_rp_r_c_is (g (l_e_st_eq_landau_n_outn x l_e_st_eq_landau_n_1 l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (l:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (g (l_e_st_eq_landau_n_outn x l_e_st_eq_landau_n_1 l)) (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t6 q x f g h pg ph l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t3 q x f g pg))))))))). Time Defined. (* constant 5494 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t8 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (l:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h l_e_st_eq_landau_n_1 l)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (l:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (g (l_e_st_eq_landau_n_outn x l_e_st_eq_landau_n_1 l)) (h (l_e_st_eq_landau_n_outn x l_e_st_eq_landau_n_1 l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t7 q x f g h pg ph l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t7 q x f h g ph pg l))))))))). Time Defined. (* constant 5495 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t9 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h l_e_st_eq_landau_n_1))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => l_andi (l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x) (forall (t:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h l_e_st_eq_landau_n_1 t) (l_e_st_eq_landau_n_satz24a x) (fun (t:l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 x) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t8 q x f g h pg ph t)))))))). Time Defined. (* constant 5496 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t10 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h l_e_st_eq_landau_n_1))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_more l_e_st_eq_landau_n_1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t9 q x f g h pg ph)))))))). Time Defined. (* constant 5497 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t19 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_not (l_e_st_eq_landau_n_more u x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_ec3e32 (l_e_st_eq_landau_n_is u x) (l_e_st_eq_landau_n_more u x) (l_e_st_eq_landau_n_less u x) (l_e_st_eq_landau_n_satz10b u x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t11 q x u l))))))))))). Time Defined. (* constant 5498 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t20 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h u)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h u) (l_e_st_eq_landau_n_more u x) p (l_e_st_eq_landau_n_rt_rp_r_c_8275_t19 q x f g h pg ph u p l))))))))))). Time Defined. (* constant 5499 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t21 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h u (l_e_st_eq_landau_n_rt_rp_r_c_8275_t12 q x u l))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_ande2 (l_e_st_eq_landau_n_lessis u x) (forall (t:l_e_st_eq_landau_n_lessis u x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h u t) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t20 q x f g h pg ph u p l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t12 q x u l))))))))))). Time Defined. (* constant 5500 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t22 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_is (q (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))) (q (h (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))) (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (h (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t21 q x f g h pg ph u p l))))))))))). Time Defined. (* constant 5501 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t23 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_is (q (h (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))) (h (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (h (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l)) (q (h (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t18 q x f h ph u l))))))))))). Time Defined. (* constant 5502 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t24 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h (l_e_st_eq_landau_n_suc u) l)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (g (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l)) (q (g (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))) (q (h (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ux q x u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t14 q x u l)))) (h (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_suc u) l)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t18 q x f g pg u l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t22 q x f g h pg ph u p l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t23 q x f g h pg ph u p l))))))))))). Time Defined. (* constant 5503 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t25 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h (l_e_st_eq_landau_n_suc u))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_andi (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) (forall (t:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h (l_e_st_eq_landau_n_suc u) t) l (fun (t:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t24 q x f g h pg ph u p t))))))))))). Time Defined. (* constant 5504 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t26 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h (l_e_st_eq_landau_n_suc u))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (l:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_suc u) x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t25 q x f g h pg ph u p l))))))))))). Time Defined. (* constant 5505 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t27 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (n:l_not (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x)), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_suc u) x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (n:l_not (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x)) => l_e_st_eq_landau_n_satz10k (l_e_st_eq_landau_n_suc u) x n)))))))))). Time Defined. (* constant 5506 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t28 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), (forall (n:l_not (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h (l_e_st_eq_landau_n_suc u))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => (fun (n:l_not (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x)) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h (l_e_st_eq_landau_n_suc u)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_suc u) x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t27 q x f g h pg ph u p n))))))))))). Time Defined. (* constant 5507 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t29 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h (l_e_st_eq_landau_n_suc u)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u) => l_imp_th1 (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h (l_e_st_eq_landau_n_suc u)) (fun (t:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t26 q x f g h pg ph u p t) (fun (t:l_not (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_suc u) x)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t28 q x f g h pg ph u p t)))))))))). Time Defined. (* constant 5508 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t30 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (u:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h u)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (u:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (v:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h v) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t10 q x f g h pg ph) (fun (v:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop5 q x f g h v) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t29 q x f g h pg ph v t)) u)))))))). Time Defined. (* constant 5509 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t31 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_is (g n) (g (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_isf (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx g n (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n)) (l_e_st_eq_landau_n_isoutinn x n))))))))). Time Defined. (* constant 5510 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t32 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (n:l_e_st_eq_landau_n_1to x), l_not (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_satz10d (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x (l_e_st_eq_landau_n_1top x n))))))))). Time Defined. (* constant 5511 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t33 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h (l_e_st_eq_landau_n_rt_rp_r_c_inn x n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (n:l_e_st_eq_landau_n_1to x) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop4 q x f g h (l_e_st_eq_landau_n_rt_rp_r_c_inn x n)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t30 q x f g h pg ph (l_e_st_eq_landau_n_rt_rp_r_c_inn x n)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t32 q x f g h pg ph n))))))))). Time Defined. (* constant 5512 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t34 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (n:l_e_st_eq_landau_n_1to x) => l_ande2 (l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x) (forall (t:l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop3 q x f g h (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) t) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t33 q x f g h pg ph n) (l_e_st_eq_landau_n_1top x n))))))))). Time Defined. (* constant 5513 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t35 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_is (h (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n))) (h n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (h n) (h (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t31 q x f h g ph pg n))))))))). Time Defined. (* constant 5514 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t36 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_is (g n) (h n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (g n) (g (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n))) (h (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_1top x n))) (h n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t31 q x f g h pg ph n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t34 q x f g h pg ph n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t35 q x f g h pg ph n))))))))). Time Defined. (* constant 5515 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t37 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h), l_e_is (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) g h))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (pg:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (ph:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => l_e_fisi (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx g h (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t36 q x f g h pg ph t)))))))). Time Defined. (* constant 5516 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_prop6 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_some (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g)))). Time Defined. (* constant 5517 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), Prop)). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => l_all (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_prop6 q x f))). Time Defined. (* constant 5518 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t38 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q l_e_st_eq_landau_n_1 f f)). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q l_e_st_eq_landau_n_1)))). Time Defined. (* constant 5519 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t39 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q l_e_st_eq_landau_n_1))), l_con)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q l_e_st_eq_landau_n_1))) => o (l_e_st_eq_landau_n_singlet_th1 n))))). Time Defined. (* constant 5520 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t40 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q l_e_st_eq_landau_n_1))), l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q l_e_st_eq_landau_n_1 n o)) (q (f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q l_e_st_eq_landau_n_1 n o))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q l_e_st_eq_landau_n_1))) => l_cone (l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q l_e_st_eq_landau_n_1 n o)) (q (f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q l_e_st_eq_landau_n_1 n o)))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t39 q f n o))))). Time Defined. (* constant 5521 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t41 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_recprop q l_e_st_eq_landau_n_1 f f)). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q l_e_st_eq_landau_n_1 f f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop2 q l_e_st_eq_landau_n_1 f f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t38 q f) (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => (fun (u:l_not (l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) t (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q l_e_st_eq_landau_n_1))) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t40 q f t u)))). Time Defined. (* constant 5522 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t42 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop6 q l_e_st_eq_landau_n_1 f)). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_somei (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q l_e_st_eq_landau_n_1 f g) f (l_e_st_eq_landau_n_rt_rp_r_c_8275_t41 q f))). Time Defined. (* constant 5523 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t43 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q l_e_st_eq_landau_n_1). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t42 q f)). Time Defined. (* constant 5524 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_xs : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), l_e_st_eq_landau_n_nat))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => l_e_st_eq_landau_n_suc x))). Time Defined. (* constant 5525 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) x (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_satz18c x)) f)))). Time Defined. (* constant 5526 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t44 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_one (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) g))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_onei (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) g) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) g) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) h) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t37 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) g h u v)))) (p (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f)))))). Time Defined. (* constant 5527 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_ind (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) g) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t44 q x p f))))). Time Defined. (* constant 5528 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t45 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_recprop q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_oneax (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) g) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t44 q x p f))))). Time Defined. (* constant 5529 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), Prop))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) n (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))))))). Time Defined. (* constant 5530 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_satz26a x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t2 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))))))). Time Defined. (* constant 5531 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), l_e_st_eq_landau_n_1to x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 q x p f n o))))))). Time Defined. (* constant 5532 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_b : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o))))))). Time Defined. (* constant 5533 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t47 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o1)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 q x p f n o1) (l_e_refis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n))))))))). Time Defined. (* constant 5534 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t48 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o1)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_isf (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t47 q x p f n o o1)))))))). Time Defined. (* constant 5535 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => q (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))))))). Time Defined. (* constant 5536 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_c : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => l_r_ite (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f))) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t48 q x p f n t u))))))). Time Defined. (* constant 5537 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t49 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => l_r_itet (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f))) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t48 q x p f n t u)) i)))))). Time Defined. (* constant 5538 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t50 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_r_itef (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f))) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t48 q x p f n t u)) o)))))). Time Defined. (* constant 5539 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_g2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f t))))). Time Defined. (* constant 5540 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t51 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_symnotis l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_ax3 x)) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) t))))). Time Defined. (* constant 5541 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t52 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t51 q x p f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t50 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t51 q x p f))))). Time Defined. (* constant 5542 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t53 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))))). Time Defined. (* constant 5543 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t54 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t51 q x p f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a x) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t51 q x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t53 q x p f))))). Time Defined. (* constant 5544 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t55 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t51 q x p f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t51 q x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t54 q x p f))))). Time Defined. (* constant 5545 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t56 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t51 q x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t52 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t55 q x p f))))). Time Defined. (* constant 5546 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t57 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t56 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t3 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t45 q x p f)))))). Time Defined. (* constant 5547 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t58 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isinoutn x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a x))))). Time Defined. (* constant 5548 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t59 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) x (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_satz18c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_satz18c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t58 q x p f))))). Time Defined. (* constant 5549 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t60 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) x (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_satz18c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t59 q x p f))))). Time Defined. (* constant 5550 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t61 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) f (l_e_st_eq_landau_n_rt_rp_r_c_8275_g2 q x p f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_1o q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t57 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t60 q x p f))))). Time Defined. (* constant 5551 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t62 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) i))))))). Time Defined. (* constant 5552 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t63 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) x))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_st_eq_landau_n_ax4 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t62 q x p f n o i)))))))). Time Defined. (* constant 5553 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t64 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 q x p f n o) x (l_e_st_eq_landau_n_lessisi3 x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t63 q x p f n o i)))))))). Time Defined. (* constant 5554 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t65 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_isf (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_symis (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t64 q x p f n o i))))))))). Time Defined. (* constant 5555 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t66 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t65 q x p f n o i) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t50 q x p f n o)))))))). Time Defined. (* constant 5556 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t67 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t66 q x p f n o i)))))))). Time Defined. (* constant 5557 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t68 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_symis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) i)))))))). Time Defined. (* constant 5558 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t69 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t68 q x p f n o i)))))))). Time Defined. (* constant 5559 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t70 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_a q x p f) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t49 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) i) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t67 q x p f n o i) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t69 q x p f n o i)))))))). Time Defined. (* constant 5560 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t71 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) x)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) => l_e_st_eq_landau_n_isoutne x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 q x p f n o) x (l_e_st_eq_landau_n_lessisi3 x) i)))))))). Time Defined. (* constant 5561 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t72 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) => l_e_st_eq_landau_n_ax2 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t71 q x p f n o o1 i))))))))). Time Defined. (* constant 5562 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t73 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t72 q x p f n o o1 i))))))))). Time Defined. (* constant 5563 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_not (l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) o1 (fun (t:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q x)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t73 q x p f n o o1 t)))))))). Time Defined. (* constant 5564 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t75 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 q x p f n o)))))))). Time Defined. (* constant 5565 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t76 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_st_eq_landau_n_ax2 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t75 q x p f n o o1)))))))). Time Defined. (* constant 5566 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t77 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)))))))). Time Defined. (* constant 5567 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t78 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_tris1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o))) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t77 q x p f n o o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t76 q x p f n o o1)))))))). Time Defined. (* constant 5568 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t79 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t46 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) o1) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o))) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t78 q x p f n o o1)))))))). Time Defined. (* constant 5569 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t80 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_isf (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t79 q x p f n o o1)))))))). Time Defined. (* constant 5570 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t81 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o))) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o))) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))))))))). Time Defined. (* constant 5571 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t82 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o))) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t76 q x p f n o o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t81 q x p f n o o1)))))))). Time Defined. (* constant 5572 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t83 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) x (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_satz18c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))) x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))) (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_satz18c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t82 q x p f n o o1)))))))). Time Defined. (* constant 5573 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t84 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) x (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) (l_e_st_eq_landau_n_satz18c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t83 q x p f n o o1)))))))). Time Defined. (* constant 5574 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t85 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o)) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t84 q x p f n o o1)))))))). Time Defined. (* constant 5575 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t86 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t50 q x p f n o)))))))). Time Defined. (* constant 5576 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t87 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t50 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t80 q x p f n o o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t4 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t45 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))))))))). Time Defined. (* constant 5577 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t88 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t86 q x p f n o o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t85 q x p f n o o1)))))))). Time Defined. (* constant 5578 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t89 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), (forall (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => (fun (o1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_b q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_n1 q x p f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t74 q x p f n o o1)))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t87 q x p f n o o1) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t88 q x p f n o o1)))))))). Time Defined. (* constant 5579 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t90 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), (forall (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (o:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f n)) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_c q x p f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o)))) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p))) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t70 q x p f n o t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_xo q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t89 q x p f n o t))))))). Time Defined. (* constant 5580 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t91 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop2 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) f (l_e_st_eq_landau_n_rt_rp_r_c_8275_g2 q x p f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_c_8275_nxs q x p f t)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t90 q x p f t u)))))). Time Defined. (* constant 5581 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t92 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_recprop q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) f (l_e_st_eq_landau_n_rt_rp_r_c_8275_g2 q x p f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) f (l_e_st_eq_landau_n_rt_rp_r_c_8275_g2 q x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop2 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) f (l_e_st_eq_landau_n_rt_rp_r_c_8275_g2 q x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t61 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t91 q x p f))))). Time Defined. (* constant 5582 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t93 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop6 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) f)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_somei (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p) f g) (l_e_st_eq_landau_n_rt_rp_r_c_8275_g2 q x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t92 q x p f))))). Time Defined. (* constant 5583 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t94 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8275_xs q x p)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t93 q x p f)))). Time Defined. (* constant 5584 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t95 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q x)). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q y) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t43 q) (fun (y:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8275_prop7 q y) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t94 q y t)) x)). Time Defined. (* constant 5585 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t96 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop6 q x f))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t95 q x f))). Time Defined. (* constant 5586 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t97 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_one (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_onei (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (h:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f h) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t37 q x f g h u v)))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t96 q x f)))). Time Defined. (* constant 5587 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz275 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_one (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t97 q x f))). Time Defined. (* constant 5588 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_recf : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_ind (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) (l_e_st_eq_landau_n_rt_rp_r_c_satz275 q x f)))). Time Defined. (* constant 5589 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz275a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_oneax (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) (l_e_st_eq_landau_n_rt_rp_r_c_satz275 q x f)))). Time Defined. (* constant 5590 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_rec : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_recf q x f n)))). Time Defined. (* constant 5591 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz275b : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_rec q x f (l_e_st_eq_landau_n_1out x)) (f (l_e_st_eq_landau_n_1out x))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t3 q x f (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f) (l_e_st_eq_landau_n_rt_rp_r_c_satz275a q x f)))). Time Defined. (* constant 5592 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_sucx : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to x), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))), l_e_st_eq_landau_n_1to x))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))) => l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x n o))))). Time Defined. (* constant 5593 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz275c : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to x), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_rec q x f (l_e_st_eq_landau_n_rt_rp_r_c_sucx q x f n o)) (q (l_e_st_eq_landau_n_rt_rp_r_c_rec q x f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_sucx q x f n o)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to x) n (l_e_st_eq_landau_n_xout x))) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t4 q x f (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f) (l_e_st_eq_landau_n_rt_rp_r_c_satz275a q x f) n o))))). Time Defined. (* constant 5594 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz275d : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), l_e_is (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) g (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t37 q x f g (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f) r (l_e_st_eq_landau_n_rt_rp_r_c_satz275a q x f)))))). Time Defined. (* constant 5595 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz275e : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_is (g n) (l_e_st_eq_landau_n_rt_rp_r_c_rec q x f n))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_c_recprop q x f g) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_fise (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx g (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f) (l_e_st_eq_landau_n_rt_rp_r_c_satz275d q x f g r) n)))))). Time Defined. (* constant 5596 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_fl : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x y l f))))). Time Defined. (* constant 5597 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_rf : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recf q x f))))). Time Defined. (* constant 5598 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x y l (l_e_st_eq_landau_n_rt_rp_r_c_8275_rf q x y l f)))))). Time Defined. (* constant 5599 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t98 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_1out y))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isinoutn y l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a y)))))). Time Defined. (* constant 5600 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t99 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1out x) (l_e_st_eq_landau_n_left1to x y l (l_e_st_eq_landau_n_1out y))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a x) (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_1out y)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_1out y)) y x (l_e_st_eq_landau_n_1top y (l_e_st_eq_landau_n_1out y)) l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t98 q x y l f)))))). Time Defined. (* constant 5601 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t100 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q y (l_e_st_eq_landau_n_rt_rp_r_c_8275_fl q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isp (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_rf q x y l f t) (f t)) (l_e_st_eq_landau_n_1out x) (l_e_st_eq_landau_n_left1to x y l (l_e_st_eq_landau_n_1out y)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t3 q x f (l_e_st_eq_landau_n_rt_rp_r_c_8275_rf q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_satz275a q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t99 q x y l f)))))). Time Defined. (* constant 5602 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t100a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_not (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_imp_th3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y) (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y)) o (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y) => l_e_tris (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_outn y (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) (l_e_st_eq_landau_n_1top y n)) (l_e_st_eq_landau_n_xout y) (l_e_st_eq_landau_n_isoutinn y n) (l_e_st_eq_landau_n_isoutni y (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) (l_e_st_eq_landau_n_1top y n) y (l_e_st_eq_landau_n_lessisi3 y) t))))))))). Time Defined. (* constant 5603 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t100b : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_ore1 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y) (l_e_st_eq_landau_n_1top y n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t100a q x y l f n o)))))))). Time Defined. (* constant 5604 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t101 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) x))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_e_st_eq_landau_n_satz16b (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t100b q x y l f n o) l))))))). Time Defined. (* constant 5605 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t102 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) x))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_ec3e31 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) x) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) x) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) x) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t101 q x y l f n o)))))))). Time Defined. (* constant 5606 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t103 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_not (l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_xout x))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_xout x)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t102 q x y l f n o) (fun (t:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_xout x)) => l_e_st_eq_landau_n_isoutne x (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y x (l_e_st_eq_landau_n_1top y n) l) x (l_e_st_eq_landau_n_lessisi3 x) t)))))))). Time Defined. (* constant 5607 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t104 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_rf q x y l f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t103 q x y l f n o))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f n) (f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t103 q x y l f n o))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t4 q x f (l_e_st_eq_landau_n_rt_rp_r_c_8275_rf q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_satz275a q x f) (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t103 q x y l f n o)))))))). Time Defined. (* constant 5608 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t105 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_left1to x y l n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) y x (l_e_st_eq_landau_n_1top y n) l)))))))). Time Defined. (* constant 5609 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t106 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn y n)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_left1to x y l n)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_e_st_eq_landau_n_ax2 (l_e_st_eq_landau_n_rt_rp_r_c_inn y n) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_left1to x y l n)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t105 q x y l f n o)))))))). Time Defined. (* constant 5610 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t107 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn y n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_e_st_eq_landau_n_isinoutn y (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn y n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q y n o)))))))). Time Defined. (* constant 5611 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t108 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_left1to x y l n))) (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_e_tris1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_left1to x y l n))) (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn y n)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t106 q x y l f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t107 q x y l f n o)))))))). Time Defined. (* constant 5612 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t109 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t103 q x y l f n o)) (l_e_st_eq_landau_n_left1to x y l (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_left1to x y l n))) (l_e_st_eq_landau_n_rt_rp_r_c_lemma275 q x (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t103 q x y l f n o)) (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o)) y x (l_e_st_eq_landau_n_1top y (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o)) l) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t108 q x y l f n o)))))))). Time Defined. (* constant 5613 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t110 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_fl q x y l f (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to y) n (l_e_st_eq_landau_n_xout y))) => l_e_isp (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8275_rf q x y l f t) (q (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f n) (f t))) (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q x (l_e_st_eq_landau_n_left1to x y l n) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t103 q x y l f n o)) (l_e_st_eq_landau_n_left1to x y l (l_e_st_eq_landau_n_rt_rp_r_c_8275_ns q y n o)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t104 q x y l f n o) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t109 q x y l f n o)))))))). Time Defined. (* constant 5614 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t111 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8275_prop2 q y (l_e_st_eq_landau_n_rt_rp_r_c_8275_fl q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (t:l_e_st_eq_landau_n_1to y) => (fun (u:l_not (l_e_is (l_e_st_eq_landau_n_1to y) t (l_e_st_eq_landau_n_xout y))) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t110 q x y l f t u))))))). Time Defined. (* constant 5615 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8275_t112 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_recprop q y (l_e_st_eq_landau_n_rt_rp_r_c_8275_fl q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop1 q y (l_e_st_eq_landau_n_rt_rp_r_c_8275_fl q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_prop2 q y (l_e_st_eq_landau_n_rt_rp_r_c_8275_fl q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f)) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t100 q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t111 q x y l f)))))). Time Defined. (* constant 5616 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz275f : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (l:l_e_st_eq_landau_n_lessis y x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x y l (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_recf q y (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x y l f))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (l:l_e_st_eq_landau_n_lessis y x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz275d q y (l_e_st_eq_landau_n_rt_rp_r_c_8275_fl q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_rfl q x y l f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t112 q x y l f)))))). Time Defined. (* constant 5617 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_xs : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_nat))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_suc x))). Time Defined. (* constant 5618 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_nat))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))). Time Defined. (* constant 5619 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))). Time Defined. (* constant 5620 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_satz18c x)))). Time Defined. (* constant 5621 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_satz4e x)))). Time Defined. (* constant 5622 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_fx : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t1 q x f) f))). Time Defined. (* constant 5623 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f) f))). Time Defined. (* constant 5624 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_f1 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t95 q x) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f)))). Time Defined. (* constant 5625 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_g2 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t95 q x) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f)))). Time Defined. (* constant 5626 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_g1 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t95 q x) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f)))). Time Defined. (* constant 5627 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_g : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) f)))). Time Defined. (* constant 5628 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8275_t49 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t95 q x) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)) (l_e_refis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)))))). Time Defined. (* constant 5629 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz275d q (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t92 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t95 q x) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f))))). Time Defined. (* constant 5630 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t6 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz275f q (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f) f))). Time Defined. (* constant 5631 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t7 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 q x f)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t6 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t5 q x f)))). Time Defined. (* constant 5632 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t8 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8276_g q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_fise (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8276_g q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t7 q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))))). Time Defined. (* constant 5633 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t9 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8276_g q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8276_g q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1 q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)))) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t8 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t4 q x f)))). Time Defined. (* constant 5634 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t10 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t2 q x f) n)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_1top x n) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t2 q x f)))))). Time Defined. (* constant 5635 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t11 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t1 q x f) n) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t2 q x f) n)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x n) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_1top x n) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t1 q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t2 q x f) n)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t2 q x f) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t2 q x f) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t10 q x f n))))). Time Defined. (* constant 5636 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t12 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f n) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x q x f n))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to x) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t1 q x f) n) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) x (l_e_st_eq_landau_n_rt_rp_r_c_8276_t2 q x f) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t11 q x f n))))). Time Defined. (* constant 5637 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t13 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x q x f)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_fisi (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x q x f) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8276_t12 q x f t)))). Time Defined. (* constant 5638 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t14 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x q x f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (u:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_recf q x u) (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t13 q x f)))). Time Defined. (* constant 5639 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t15 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x q x f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz275d q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8275_t45 q x (l_e_st_eq_landau_n_rt_rp_r_c_8275_t95 q x) (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1 q x f))))). Time Defined. (* constant 5640 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t16 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_f1x q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t15 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t14 q x f)))). Time Defined. (* constant 5641 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t17 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f) (l_e_st_eq_landau_n_xout x))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_fise (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t16 q x f) (l_e_st_eq_landau_n_xout x)))). Time Defined. (* constant 5642 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t18 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))))). Time Defined. (* constant 5643 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t19 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (l_e_st_eq_landau_n_satz4a x) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t18 q x f)))). Time Defined. (* constant 5644 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t20 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t19 q x f)))). Time Defined. (* constant 5645 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t21 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isp1 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) f t) (q (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (f t))) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t3 q x f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_xs q x f))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t9 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t20 q x f)))). Time Defined. (* constant 5646 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8276_t22 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f) (l_e_st_eq_landau_n_xout x)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8276_x1 q x f)))) (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f) (l_e_st_eq_landau_n_xout x)) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t17 q x f)))). Time Defined. (* constant 5647 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz276 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (q (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)) f) (l_e_st_eq_landau_n_xout x)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8276_g1x q x f (l_e_st_eq_landau_n_xout x)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_recf q x (l_e_st_eq_landau_n_rt_rp_r_c_8276_fx q x f) (l_e_st_eq_landau_n_xout x)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t21 q x f) (l_e_st_eq_landau_n_rt_rp_r_c_8276_t22 q x f)))). Time Defined. (* constant 5648 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_smpr : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_cx))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_rec q x f (l_e_st_eq_landau_n_xout x)))). Time Defined. (* constant 5649 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_sum : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_cx)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl t u)) x f)). Time Defined. (* constant 5650 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_prod : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_cx)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) x f)). Time Defined. (* constant 5651 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8277_t1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_lessisi3 l_e_st_eq_landau_n_1) (l_e_refis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1))). Time Defined. (* constant 5652 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz277 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isp (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_rec q l_e_st_eq_landau_n_1 f t) (f t)) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_satz275b q l_e_st_eq_landau_n_1 f) (l_e_st_eq_landau_n_rt_rp_r_c_8277_t1 q f))). Time Defined. (* constant 5653 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz278 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)) f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz276 q x f))). Time Defined. (* constant 5654 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), l_e_st_eq_landau_n_lessis y x))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => l_e_st_eq_landau_n_lessisi2 y x i))))). Time Defined. (* constant 5655 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v8_f0 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), (forall (t:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x y (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i) f))))). Time Defined. (* constant 5656 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v8_t2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i) (l_e_st_eq_landau_n_xout y))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_v8_f0 q x f y i))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => l_e_fise (l_e_st_eq_landau_n_1to y) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x y (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f)) (l_e_st_eq_landau_n_rt_rp_r_c_recf q y (l_e_st_eq_landau_n_rt_rp_r_c_v8_f0 q x f y i)) (l_e_st_eq_landau_n_rt_rp_r_c_satz275f q x y (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i) f) (l_e_st_eq_landau_n_xout y)))))). Time Defined. (* constant 5657 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v8_t3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_xout y)) x))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => l_e_tris1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_xout y)) x y (l_e_st_eq_landau_n_isinoutn y y (l_e_st_eq_landau_n_lessisi3 y)) i))))). Time Defined. (* constant 5658 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v8_t4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i) (l_e_st_eq_landau_n_xout y)) (l_e_st_eq_landau_n_xout x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_xout y)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn y (l_e_st_eq_landau_n_xout y)) y x (l_e_st_eq_landau_n_1top y (l_e_st_eq_landau_n_xout y)) (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i)) x (l_e_st_eq_landau_n_lessisi3 x) (l_e_st_eq_landau_n_rt_rp_r_c_v8_t3 q x f y i)))))). Time Defined. (* constant 5659 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v8_t5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i) (l_e_st_eq_landau_n_xout y))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => l_e_isf (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f) (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i) (l_e_st_eq_landau_n_xout y)) (l_e_st_eq_landau_n_xout x) (l_e_st_eq_landau_n_rt_rp_r_c_v8_t4 q x f y i)))))). Time Defined. (* constant 5660 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_issmpr : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (y:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is y x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x y (l_e_st_eq_landau_n_lessisi2 y x i) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is y x) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_v8_f0 q x f y i)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x f) (l_e_st_eq_landau_n_rt_rp_r_c_recf q x f (l_e_st_eq_landau_n_left1to x y (l_e_st_eq_landau_n_rt_rp_r_c_v8_t1 q x f y i) (l_e_st_eq_landau_n_xout y))) (l_e_st_eq_landau_n_rt_rp_r_c_v8_t2 q x f y i) (l_e_st_eq_landau_n_rt_rp_r_c_v8_t5 q x f y i)))))). Time Defined. (* constant 5661 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_xr : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt x)). Time Defined. (* constant 5662 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), Prop)). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum x (fun (t:l_e_st_eq_landau_n_1to x) => z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 5663 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t1 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum l_e_st_eq_landau_n_1 (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => z)) z). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl t u)) (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => z)). Time Defined. (* constant 5664 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t2 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z l_e_st_eq_landau_n_1). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_sum l_e_st_eq_landau_n_1 (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => z)) z (l_e_st_eq_landau_n_rt_rp_r_c_ts z l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t1 z) (l_e_st_eq_landau_n_rt_rp_r_c_satz222a z)). Time Defined. (* constant 5665 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t3 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (fun (t:l_e_st_eq_landau_n_1to x) => z)) z)))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl t u)) x (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => z)))). Time Defined. (* constant 5666 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t4 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (fun (t:l_e_st_eq_landau_n_1to x) => z)) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z l_e_st_eq_landau_n_rt_rp_r_c_1c))))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_st_eq_landau_n_rt_rp_r_c_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_sum x (fun (t:l_e_st_eq_landau_n_1to x) => z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0)) z (l_e_st_eq_landau_n_rt_rp_r_c_ts z l_e_st_eq_landau_n_rt_rp_r_c_1c) p (l_e_st_eq_landau_n_rt_rp_r_c_satz222a z)))). Time Defined. (* constant 5667 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t5 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_1c))))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_st_eq_landau_n_rt_rp_r_c_distpt2 z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_1c))). Time Defined. (* constant 5668 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t6 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_st_eq_landau_n_rt_rp_r_c_plis12a (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 5669 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t7 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satzr155b x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 5670 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t8 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t6 z x p) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t7 z x p)))). Time Defined. (* constant 5671 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t9 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_0) z (l_e_st_eq_landau_n_rt_rp_r_c_8279_t8 z x p)))). Time Defined. (* constant 5672 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t10 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => z)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (fun (t:l_e_st_eq_landau_n_1to x) => z)) z) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8279_xr z x) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t3 z x p) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t4 z x p) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t5 z x p) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t9 z x p)))). Time Defined. (* constant 5673 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8279_t11 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x), l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z (l_e_st_eq_landau_n_suc x)))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z x) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z t) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t10 z x p) (l_e_st_eq_landau_n_satz4a x)))). Time Defined. (* constant 5674 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz279 : (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum x (fun (t:l_e_st_eq_landau_n_1to x) => z)) (l_e_st_eq_landau_n_rt_rp_r_c_ts z (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_rlofnt x) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z t) (l_e_st_eq_landau_n_rt_rp_r_c_8279_t2 z) (fun (u:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8279_prop1 z u) => l_e_st_eq_landau_n_rt_rp_r_c_8279_t11 z u t)) x)). Time Defined. (* constant 5675 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_t1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_2)). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_satz18a l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1))). Time Defined. (* constant 5676 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_t1 q f) f)). Time Defined. (* constant 5677 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_t2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 f) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f)) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 q l_e_st_eq_landau_n_1 f)). Time Defined. (* constant 5678 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_t3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f)) (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 q (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f))). Time Defined. (* constant 5679 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_t4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isinoutn l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_lessisi3 l_e_st_eq_landau_n_1))). Time Defined. (* constant 5680 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_t5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2) (l_e_st_eq_landau_n_left1to l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_t1 q f) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a l_e_st_eq_landau_n_2) (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_1top l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8280_t1 q f)) (l_e_st_eq_landau_n_rt_rp_r_c_8280_t4 q f))). Time Defined. (* constant 5681 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_t6 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2) (l_e_st_eq_landau_n_left1to l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_t1 q f) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8280_t5 q f))). Time Defined. (* constant 5682 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_t7 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f)) (f (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f)) (f (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8280_t3 q f) (l_e_st_eq_landau_n_rt_rp_r_c_8280_t6 q f))). Time Defined. (* constant 5683 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8280_t8 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f)) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))) (q (f (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f)) (f (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8280_t7 q f))). Time Defined. (* constant 5684 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz280 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 f) (q (f (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 f) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8280_f1 q f)) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))) (q (f (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))) (l_e_st_eq_landau_n_rt_rp_r_c_8280_t2 q f) (l_e_st_eq_landau_n_rt_rp_r_c_8280_t8 q f))). Time Defined. (* constant 5685 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_assoc : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), Prop). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q x y) z) (q x (q y z)))))). Time Defined. (* constant 5686 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_assocp1 : l_e_st_eq_landau_n_rt_rp_r_c_assoc (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl x y)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_asspl1 x y z))). Time Defined. (* constant 5687 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_assocts : l_e_st_eq_landau_n_rt_rp_r_c_assoc (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts x y)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_assts1 x y z))). Time Defined. (* constant 5688 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_assq1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q z u) v) (q z (q u v))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => a z u v))))). Time Defined. (* constant 5689 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_assq2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q z (q u v)) (q (q z u) v)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (q (q z u) v) (q z (q u v)) (l_e_st_eq_landau_n_rt_rp_r_c_assq1 q a z u v)))))). Time Defined. (* constant 5690 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_pl x y))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_satz18a x y))))). Time Defined. (* constant 5691 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x y) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_t1 q a x y) f))))). Time Defined. (* constant 5692 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_right l_e_st_eq_landau_n_rt_rp_r_c_cx x y f))))). Time Defined. (* constant 5693 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_prop1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl x y) f) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y f)))))))). Time Defined. (* constant 5694 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), Prop)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_all (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8281_prop1 q a x y u))))). Time Defined. (* constant 5695 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f0) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x l_e_st_eq_landau_n_1 f0)) (f0 (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 q x f0)))). Time Defined. (* constant 5696 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 q (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0))))). Time Defined. (* constant 5697 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isinoutn l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_lessisi3 l_e_st_eq_landau_n_1))))). Time Defined. (* constant 5698 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_ispl2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_t4 q a x f0))))). Time Defined. (* constant 5699 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t6 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_right1to x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_1top l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t5 q a x f0))))). Time Defined. (* constant 5700 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t7 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (f0 (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_c_cx f0 (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_right1to x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t6 q a x f0))))). Time Defined. (* constant 5701 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t8 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (f0 (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (f0 (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t7 q a x f0) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t3 q a x f0))))). Time Defined. (* constant 5702 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t9 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x l_e_st_eq_landau_n_1 f0)) (f0 (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x l_e_st_eq_landau_n_1 f0)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x l_e_st_eq_landau_n_1 f0))) (f0 (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t8 q a x f0))))). Time Defined. (* constant 5703 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t10 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8281_prop1 q a x l_e_st_eq_landau_n_1 f0)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f0) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x l_e_st_eq_landau_n_1 f0)) (f0 (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x l_e_st_eq_landau_n_1 f0)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x l_e_st_eq_landau_n_1 f0))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t2 q a x f0) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t9 q a x f0))))). Time Defined. (* constant 5704 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t11 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x l_e_st_eq_landau_n_1))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8281_t10 q a x u)))). Time Defined. (* constant 5705 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl y l_e_st_eq_landau_n_1)))). Time Defined. (* constant 5706 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl x y)))). Time Defined. (* constant 5707 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))))). Time Defined. (* constant 5708 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) l_e_st_eq_landau_n_1)))). Time Defined. (* constant 5709 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_asspl1 x y l_e_st_eq_landau_n_1))))). Time Defined. (* constant 5710 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t13 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))))))). Time Defined. (* constant 5711 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t14 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_asspl2 x y l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t13 q a x y p f))))))). Time Defined. (* constant 5712 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t15 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t14 q a x y p f))))))). Time Defined. (* constant 5713 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t16 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t15 q a x y p f))))))). Time Defined. (* constant 5714 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_fr : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) f)))))). Time Defined. (* constant 5715 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t17 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f)) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz275f q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) f)))))). Time Defined. (* constant 5716 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t18 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_fise (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f)) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t17 q a x y p f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))))))). Time Defined. (* constant 5717 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t19 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_satz18a (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) l_e_st_eq_landau_n_1))))))). Time Defined. (* constant 5718 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_frr : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t19 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f))))))). Time Defined. (* constant 5719 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t20 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f))))))). Time Defined. (* constant 5720 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t21 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q u (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))) (p (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))))))). Time Defined. (* constant 5721 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t22 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_assq1 q a (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))))))). Time Defined. (* constant 5722 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t23 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 y (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_satz18a y l_e_st_eq_landau_n_1))))))). Time Defined. (* constant 5723 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_fy : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) y (l_e_st_eq_landau_n_rt_rp_r_c_8281_t23 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f))))))). Time Defined. (* constant 5724 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t24 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f))))))). Time Defined. (* constant 5725 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t25 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))))). Time Defined. (* constant 5726 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t26 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_t25 q a x y p f))))))). Time Defined. (* constant 5727 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t27 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t14 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t26 q a x y p f))))))). Time Defined. (* constant 5728 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t28 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_right1to x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) x (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t27 q a x y p f))))))). Time Defined. (* constant 5729 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t29 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_right1to x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t28 q a x y p f))))))). Time Defined. (* constant 5730 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t30 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t29 q a x y p f))))))). Time Defined. (* constant 5731 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_nat))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_rt_rp_r_c_inn y n))))))). Time Defined. (* constant 5732 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) y (l_e_st_eq_landau_n_rt_rp_r_c_8281_t23 q a x y p f) n))))))). Time Defined. (* constant 5733 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t31 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n) y (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_1top y n) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t23 q a x y p f))))))))). Time Defined. (* constant 5734 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t32 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n)) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_t31 q a x y p f n)))))))). Time Defined. (* constant 5735 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_nxpy : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_right1to x y n))))))). Time Defined. (* constant 5736 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t19 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxpy q a x y p f n)))))))). Time Defined. (* constant 5737 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t33 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxpy q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxpy q a x y p f n)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxpy q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxpy q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t19 q a x y p f))))))))). Time Defined. (* constant 5738 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t34 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxpy q a x y p f n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n)) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n) y x (l_e_st_eq_landau_n_1top y n))))))))). Time Defined. (* constant 5739 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t35 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxpy q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t34 q a x y p f n) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t33 q a x y p f n)))))))). Time Defined. (* constant 5740 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t36 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n))) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_tris1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n))) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n)) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_8281_n0 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t32 q a x y p f n) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t35 q a x y p f n)))))))). Time Defined. (* constant 5741 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t37 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) (l_e_st_eq_landau_n_right1to x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_pl x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n))) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) x (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n))) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t36 q a x y p f n)))))))). Time Defined. (* constant 5742 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t38 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f n) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f) n)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to y) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_right1to x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nyp1 q a x y p f n)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_nxyp1 q a x y p f n)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t37 q a x y p f n)))))))). Time Defined. (* constant 5743 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t39 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_fisi (l_e_st_eq_landau_n_1to y) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)) (fun (u:l_e_st_eq_landau_n_1to y) => l_e_st_eq_landau_n_rt_rp_r_c_8281_t38 q a x y p f u))))))). Time Defined. (* constant 5744 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t40 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:(forall (u:l_e_st_eq_landau_n_1to y), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q y t) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t39 q a x y p f))))))). Time Defined. (* constant 5745 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t41 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f)) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t40 q a x y p f)))))))). Time Defined. (* constant 5746 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t41a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t30 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t41 q a x y p f))))))). Time Defined. (* constant 5747 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t42 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_fy q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t41a q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t24 q a x y p f))))))). Time Defined. (* constant 5748 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t43 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t42 q a x y p f))))))). Time Defined. (* constant 5749 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_m0 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (m:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (m:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_inn x m))))))). Time Defined. (* constant 5750 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_mxpy : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (m:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (m:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_t1 q a x y) m))))))). Time Defined. (* constant 5751 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (m:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (m:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t19 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxpy q a x y p f m)))))))). Time Defined. (* constant 5752 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t44 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (m:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxpy q a x y p f m)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 q a x y p f m))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (m:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxpy q a x y p f m)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxpy q a x y p f m)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxpy q a x y p f m)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t19 q a x y p f))))))))). Time Defined. (* constant 5753 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t45 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (m:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_m0 q a x y p f m) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxpy q a x y p f m))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (m:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_m0 q a x y p f m) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_8281_m0 q a x y p f m) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_1top x m) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t1 q a x y))))))))). Time Defined. (* constant 5754 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t46 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (m:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_m0 q a x y p f m) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 q a x y p f m))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (m:l_e_st_eq_landau_n_1to x) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_8281_m0 q a x y p f m) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxpy q a x y p f m)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 q a x y p f m)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t45 q a x y p f m) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t44 q a x y p f m)))))))). Time Defined. (* constant 5755 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t47 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (m:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_t1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)) m) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 q a x y p f m))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (m:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_m0 q a x y p f m) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_8281_m0 q a x y p f m) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_1top x m) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 q a x y p f m)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 q a x y p f m)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 q a x y p f m)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t46 q a x y p f m)))))))). Time Defined. (* constant 5756 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t48 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (m:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f m) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f) m)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (m:l_e_st_eq_landau_n_1to x) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) x (l_e_st_eq_landau_n_rt_rp_r_c_8281_t1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)) m) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_mxyp1 q a x y p f m)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t47 q a x y p f m)))))))). Time Defined. (* constant 5757 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t49 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_fisi (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)) (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8281_t48 q a x y p f u))))))). Time Defined. (* constant 5758 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t50 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:(forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q x t) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t49 q a x y p f))))))). Time Defined. (* constant 5759 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t51 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t50 q a x y p f)))))))). Time Defined. (* constant 5760 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t52 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t12 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_recf q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t16 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t18 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t20 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t21 q a x y p f))))))). Time Defined. (* constant 5761 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t53 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8281_prop1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y) f) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f)))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_fr q a x y p f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8281_xyp1 q a x y))))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x y (l_e_st_eq_landau_n_rt_rp_r_c_8281_frr q a x y p f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8281_f1 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_f2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) f))) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t52 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t22 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t43 q a x y p f) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t51 q a x y p f))))))). Time Defined. (* constant 5762 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t54 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8281_xpy1 q a x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8281_t53 q a x y p u)))))). Time Defined. (* constant 5763 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t55 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y), l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x (l_e_st_eq_landau_n_suc y)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x t) (l_e_st_eq_landau_n_rt_rp_r_c_8281_yp1 q a x y) (l_e_st_eq_landau_n_suc y) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t54 q a x y p) (l_e_st_eq_landau_n_satz4a y)))))). Time Defined. (* constant 5764 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8281_t56 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x y)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x t) (l_e_st_eq_landau_n_rt_rp_r_c_8281_t11 q a x) (fun (z:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8281_prop2 q a x z) => l_e_st_eq_landau_n_rt_rp_r_c_8281_t55 q a x z t)) y)))). Time Defined. (* constant 5765 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz281 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (x:l_e_st_eq_landau_n_nat), (forall (y:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl x y) f) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x y) x (l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_pl x y) (l_e_st_eq_landau_n_satz18a x y)) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q y (l_e_st_eq_landau_n_right l_e_st_eq_landau_n_rt_rp_r_c_cx x y f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (y:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x y)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8281_t56 q a x y f))))). Time Defined. (* constant 5766 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_commut : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), Prop). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q x y) (q y x)))). Time Defined. (* constant 5767 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_commutpl : l_e_st_eq_landau_n_rt_rp_r_c_commut (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl x y)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_compl x y)). Time Defined. (* constant 5768 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_commutts : l_e_st_eq_landau_n_rt_rp_r_c_commut (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts x y)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_comts x y)). Time Defined. (* constant 5769 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_comq : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q z u) (q u z))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => c z u)))). Time Defined. (* constant 5770 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_prop1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (fun (t:l_e_st_eq_landau_n_1to x) => q (f t) (g t))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x f) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x g)))))))). Time Defined. (* constant 5771 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), Prop)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => l_all (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (u:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_all (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (v:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8282_prop1 q a c x u v)))))). Time Defined. (* constant 5772 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => q (f0 t) (g0 t))) (q (f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 q (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => q (f0 t) (g0 t))))))). Time Defined. (* constant 5773 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 g0)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 g0) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_satz277 q g0)))))). Time Defined. (* constant 5774 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (q (f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0) (f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_satz277 q f0)))))). Time Defined. (* constant 5775 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 g0)) (q (f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 g0)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (q (f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t2 q a c f0 g0) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t3 q a c f0 g0)))))). Time Defined. (* constant 5776 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8282_prop1 q a c l_e_st_eq_landau_n_1 f0 g0))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (f0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g0:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => q (f0 t) (g0 t))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f0) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 g0)) (q (f0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (g0 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t1 q a c f0 g0) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t4 q a c f0 g0)))))). Time Defined. (* constant 5777 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t6 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c l_e_st_eq_landau_n_1))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (v:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8282_t5 q a c u v))))). Time Defined. (* constant 5778 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))). Time Defined. (* constant 5779 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t7 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q u v) (q w z)) (q (q (q u v) w) z)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_assq2 q a (q u v) w z))))))). Time Defined. (* constant 5780 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t8 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q u v) w) (q w (q u v))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_comq q c (q u v) w))))))). Time Defined. (* constant 5781 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t9 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q w (q u v)) (q (q w u) v)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_assq2 q a w u v))))))). Time Defined. (* constant 5782 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t10 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q u v) w) (q (q w u) v)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (q (q u v) w) (q w (q u v)) (q (q w u) v) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t8 q a c u v w z) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t9 q a c u v w z)))))))). Time Defined. (* constant 5783 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t11 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q (q u v) w) z) (q (q (q w u) v) z)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t z) (q (q u v) w) (q (q w u) v) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t10 q a c u v w z)))))))). Time Defined. (* constant 5784 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t12 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q (q w u) v) z) (q (q w u) (q v z))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_assq1 q a (q w u) v z))))))). Time Defined. (* constant 5785 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t13 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q w u) (q u w)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_comq q c w u))))))). Time Defined. (* constant 5786 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t14 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q w u) (q v z)) (q (q u w) (q v z))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (q v z)) (q w u) (q u w) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t13 q a c u v w z)))))))). Time Defined. (* constant 5787 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t15 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (z:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q u v) (q w z)) (q (q u w) (q v z))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (z:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (q (q u v) (q w z)) (q (q (q u v) w) z) (q (q (q w u) v) z) (q (q w u) (q v z)) (q (q u w) (q v z)) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t7 q a c u v w z) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t11 q a c u v w z) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t12 q a c u v w z) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t14 q a c u v w z)))))))). Time Defined. (* constant 5788 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t16 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))))))). Time Defined. (* constant 5789 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8282_t16 q a c x p f g) f)))))))). Time Defined. (* constant 5790 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8282_t16 q a c x p f g) g)))))))). Time Defined. (* constant 5791 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_h : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)) => q (f t) (g t))))))))). Time Defined. (* constant 5792 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_shx : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8282_t16 q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_h q a c x p f g))))))))). Time Defined. (* constant 5793 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t17 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8282_h q a c x p f g)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_shx q a c x p f g) (q (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 q x (l_e_st_eq_landau_n_rt_rp_r_c_8282_h q a c x p f g)))))))). Time Defined. (* constant 5794 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t18 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8282_shx q a c x p f g) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => p (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8282_t16 q a c x p f g) f) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8282_t16 q a c x p f g) g)))))))). Time Defined. (* constant 5795 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t19 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_shx q a c x p f g) (q (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g)) (q (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (q (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (l_e_st_eq_landau_n_rt_rp_r_c_8282_shx q a c x p f g) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g)) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t18 q a c x p f g)))))))). Time Defined. (* constant 5796 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t20 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g)) (q (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8282_t15 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))))))))). Time Defined. (* constant 5797 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t21 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_satz278 q x f)))))))). Time Defined. (* constant 5798 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t22 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t21 q a c x p f g)))))))). Time Defined. (* constant 5799 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t23 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) g)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) g) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_satz278 q x g)))))))). Time Defined. (* constant 5800 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t24 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) g))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t23 q a c x p f g)))))))). Time Defined. (* constant 5801 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t25 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8282_h q a c x p f g)) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8282_h q a c x p f g)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_shx q a c x p f g) (q (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g)) (q (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t17 q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t19 q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t20 q a c x p f g)))))))). Time Defined. (* constant 5802 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t26 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8282_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f g))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8282_h q a c x p f g)) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sfx q a c x p f g) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f) (q (l_e_st_eq_landau_n_rt_rp_r_c_8282_sgx q a c x p f g) (g (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x))))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) g)) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t25 q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t22 q a c x p f g) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t24 q a c x p f g)))))))). Time Defined. (* constant 5803 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t27 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (v:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8282_t26 q a c x p u v))))))). Time Defined. (* constant 5804 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t28 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x), l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c (l_e_st_eq_landau_n_suc x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c t) (l_e_st_eq_landau_n_rt_rp_r_c_8282_xp1 q a c x) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t27 q a c x p) (l_e_st_eq_landau_n_satz4a x)))))). Time Defined. (* constant 5805 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8282_t29 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c t) (l_e_st_eq_landau_n_rt_rp_r_c_8282_t6 q a c) (fun (y:l_e_st_eq_landau_n_nat) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8282_prop2 q a c y) => l_e_st_eq_landau_n_rt_rp_r_c_8282_t28 q a c y t)) x)))). Time Defined. (* constant 5806 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz282 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (fun (t:l_e_st_eq_landau_n_1to x) => q (f t) (g t))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x f) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x g)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (g:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8282_t29 q a c x f g)))))). Time Defined. (* constant 5807 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (s:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (t:l_e_st_eq_landau_n_1to x) => f (s t)))))))). Time Defined. (* constant 5808 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (s:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c x s f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x f))))))). Time Defined. (* constant 5809 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), Prop)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => l_all (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x) (fun (u:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)) => l_all (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (v:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_imp (l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c x u v))))))). Time Defined. (* constant 5810 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (s (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (s (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_1o (l_e_st_eq_landau_n_singlet_th1 (s (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_singlet_th1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))). Time Defined. (* constant 5811 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c l_e_st_eq_landau_n_1 s f)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c l_e_st_eq_landau_n_1 s f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c l_e_st_eq_landau_n_1 s f)))))). Time Defined. (* constant 5812 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c l_e_st_eq_landau_n_1 s f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_isf (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_rt_rp_r_c_cx f (s (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t1 q a c s f)))))). Time Defined. (* constant 5813 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_satz277 q f)))))). Time Defined. (* constant 5814 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)), (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c l_e_st_eq_landau_n_1 s f))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c l_e_st_eq_landau_n_1 s f)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c l_e_st_eq_landau_n_1 s f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t2 q a c s f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t3 q a c s f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t4 q a c s f)))))). Time Defined. (* constant 5815 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t6 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c l_e_st_eq_landau_n_1))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1)) => (fun (v:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (w:l_e_bijective (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) u) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t5 q a c u v)))))). Time Defined. (* constant 5816 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))). Time Defined. (* constant 5817 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1))))). Time Defined. (* constant 5818 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t8 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), l_e_injective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => l_ande1 (l_e_injective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) (l_e_surjective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) b)))))))). Time Defined. (* constant 5819 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x) u)))))))))). Time Defined. (* constant 5820 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)))))))))))). Time Defined. (* constant 5821 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t9 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_tr3is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u))) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i) (l_e_symis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) case1)))))))))))). Time Defined. (* constant 5822 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t10 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_isfe (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_t8 q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t9 q a c x p s f b case1 u i)))))))))))). Time Defined. (* constant 5823 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t11 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t10 q a c x p s f b case1 u i)))))))))))). Time Defined. (* constant 5824 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t12 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_satz16a (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x u) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))))))))))). Time Defined. (* constant 5825 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t13 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_con))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_ec3e31 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t12 q a c x p s f b case1 u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t11 q a c x p s f b case1 u i)))))))))))). Time Defined. (* constant 5826 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t14 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_ore1 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u))) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t13 q a c x p s f b case1 u t))))))))))). Time Defined. (* constant 5827 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t15 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_satz26 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t14 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5828 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t15 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5829 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 t)))))))))). Time Defined. (* constant 5830 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t16 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 v)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 v))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 v)) => l_e_st_eq_landau_n_isoutne x (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t15 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 v) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t15 q a c x p s f b case1 v) i)))))))))))). Time Defined. (* constant 5831 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t17 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 v)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 v))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 v)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t8 q a c x p s f b (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 v) (l_e_st_eq_landau_n_isinne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 v)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t16 q a c x p s f b case1 u v i)))))))))))))). Time Defined. (* constant 5832 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t18 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 v)), l_e_is (l_e_st_eq_landau_n_1to x) u v)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 v)) => l_e_st_eq_landau_n_thleft1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x) u v (l_e_st_eq_landau_n_rt_rp_r_c_8283_t17 q a c x p s f b case1 u v i))))))))))))). Time Defined. (* constant 5833 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_invf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s b (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5834 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5835 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t19 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i)))))))))))). Time Defined. (* constant 5836 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t20 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_tr3is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_thinvf2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s b (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t19 q a c x p s f b case1 u i)) case1))))))))))). Time Defined. (* constant 5837 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t21 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t20 q a c x p s f b case1 u i)))))))))))). Time Defined. (* constant 5838 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t22 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_satz16a (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x u) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))))))))))). Time Defined. (* constant 5839 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t23 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_con))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_ec3e31 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t22 q a c x p s f b case1 u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t21 q a c x p s f b case1 u i)))))))))))). Time Defined. (* constant 5840 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t24 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_ore1 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u)) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t23 q a c x p s f b case1 u t))))))))))). Time Defined. (* constant 5841 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t25 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_satz26 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t24 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5842 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t25 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5843 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t26 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t25 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5844 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t27 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t26 q a c x p s f b case1 u)))))))))))). Time Defined. (* constant 5845 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t28 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u))) (l_e_thinvf2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s b (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t27 q a c x p s f b case1 u)))))))))))). Time Defined. (* constant 5846 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t29 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x))) (l_e_st_eq_landau_n_isinni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t28 q a c x p s f b case1 u)))))))))))). Time Defined. (* constant 5847 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t30 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to x) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris (l_e_st_eq_landau_n_1to x) u (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_1top x u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_isoutinn x u) (l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_1top x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t15 q a c x p s f b case1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t29 q a c x p s f b case1 u)))))))))))). Time Defined. (* constant 5848 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t31 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_image (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1) u)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_somei (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_is (l_e_st_eq_landau_n_1to x) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 t)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w2 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t30 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5849 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t32 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_andi (l_e_injective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1)) (l_e_surjective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1)) (fun (t:l_e_st_eq_landau_n_1to x) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1 t) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1 u)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t18 q a c x p s f b case1 t u v))) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t31 q a c x p s f b case1 t)))))))))). Time Defined. (* constant 5850 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x) f))))))))). Time Defined. (* constant 5851 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t33 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c x (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t32 q a c x p s f b case1)))))))))). Time Defined. (* constant 5852 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))). Time Defined. (* constant 5853 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c x (l_e_st_eq_landau_n_rt_rp_r_c_8283_s01 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1)))))))))). Time Defined. (* constant 5854 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t33a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)))))))))))). Time Defined. (* constant 5855 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t34 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t15 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5856 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t35 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)))) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u))) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u)) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t34 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5857 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t36 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)))) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t33a q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t35 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5858 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t37 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g2 q a c x p s f b case1 u))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t7 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w1 q a c x p s f b case1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t36 q a c x p s f b case1 u))))))))))). Time Defined. (* constant 5859 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t38 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g2 q a c x p s f b case1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_fisi (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g2 q a c x p s f b case1) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t37 q a c x p s f b case1 t)))))))))). Time Defined. (* constant 5860 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t39 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g2 q a c x p s f b case1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (u:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g2 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t38 q a c x p s f b case1)))))))))). Time Defined. (* constant 5861 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t40 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g2 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t39 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t33 q a c x p s f b case1)))))))))). Time Defined. (* constant 5862 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t41 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) case1))))))))). Time Defined. (* constant 5863 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t42 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))). Time Defined. (* constant 5864 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t43 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t41 q a c x p s f b case1)))))))))). Time Defined. (* constant 5865 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t44 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t40 q a c x p s f b case1)))))))))). Time Defined. (* constant 5866 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t45 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_satz278 q x f)))))))))). Time Defined. (* constant 5867 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t46 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case1:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g1 q a c x p s f b case1)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f01 q a c x p s f b case1)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t42 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t43 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t44 q a c x p s f b case1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t45 q a c x p s f b case1)))))))))). Time Defined. (* constant 5868 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_1px : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), l_e_st_eq_landau_n_nat)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 x)))))))). Time Defined. (* constant 5869 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x u)))))))))). Time Defined. (* constant 5870 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_compl l_e_st_eq_landau_n_1 x)))))))))). Time Defined. (* constant 5871 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) s))))))))). Time Defined. (* constant 5872 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)))))))))))). Time Defined. (* constant 5873 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t48 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_tr3is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i) (l_e_symis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) case2)))))))))))). Time Defined. (* constant 5874 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t49 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_isfe (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_t8 q a c x p s f b) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t48 q a c x p s f b case2 u i)))))))))))). Time Defined. (* constant 5875 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t50 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t49 q a c x p s f b case2 u i)))))))))))). Time Defined. (* constant 5876 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t51 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_1top x u))))))))))))). Time Defined. (* constant 5877 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t52 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t51 q a c x p s f b case2 u i) (l_e_st_eq_landau_n_satz18 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u))))))))))))). Time Defined. (* constant 5878 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t53 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_con))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_ec3e21 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t52 q a c x p s f b case2 u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t50 q a c x p s f b case2 u i)))))))))))). Time Defined. (* constant 5879 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t54 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_ore1 (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz24 (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u)) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t53 q a c x p s f b case2 u t))))))))))). Time Defined. (* constant 5880 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_mn (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t54 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5881 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t54a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_islessis1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_mn_th1c (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t54 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))))))))))))). Time Defined. (* constant 5882 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t55 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_or_th9 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) x) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t54a q a c x p s f b case2 u) (fun (t:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_satz20c (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) x l_e_st_eq_landau_n_1 t) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_satz20b (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) x l_e_st_eq_landau_n_1 t))))))))))). Time Defined. (* constant 5883 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t55 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5884 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 t)))))))))). Time Defined. (* constant 5885 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t56 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 v))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)) => l_e_st_eq_landau_n_isoutne x (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t55 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 v) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t55 q a c x p s f b case2 v) i)))))))))))). Time Defined. (* constant 5886 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t57 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 v))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 v) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 v) (l_e_st_eq_landau_n_mn_th1c (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t54 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 v) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t56 q a c x p s f b case2 u v i)) (l_e_st_eq_landau_n_mn_th1d (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 v) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t54 q a c x p s f b case2 v)))))))))))))). Time Defined. (* constant 5887 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t58 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 v)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)) => l_e_st_eq_landau_n_isinne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 v)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t57 q a c x p s f b case2 u v i))))))))))))). Time Defined. (* constant 5888 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t59 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 v)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)) => l_e_isfe (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_t8 q a c x p s f b) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 v)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t58 q a c x p s f b case2 u v i))))))))))))). Time Defined. (* constant 5889 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t60 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 v))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)) => l_e_st_eq_landau_n_thleft1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 v) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t59 q a c x p s f b case2 u v i))))))))))))). Time Defined. (* constant 5890 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t61 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (v:l_e_st_eq_landau_n_1to x), (forall (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)), l_e_is (l_e_st_eq_landau_n_1to x) u v)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 v)) => l_e_st_eq_landau_n_thright1 l_e_st_eq_landau_n_1 x u v (l_e_st_eq_landau_n_rt_rp_r_c_8283_t60 q a c x p s f b case2 u v i))))))))))))). Time Defined. (* constant 5891 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s04 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_invf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s b)))))))))). Time Defined. (* constant 5892 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8283_s04 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5893 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5894 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t62 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i)))))))))))). Time Defined. (* constant 5895 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t63 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_tr3is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_thinvf2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s b (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t62 q a c x p s f b case2 u i)) case2))))))))))). Time Defined. (* constant 5896 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t64 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t63 q a c x p s f b case2 u i)))))))))))). Time Defined. (* constant 5897 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t65 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) l_e_st_eq_landau_n_1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_1top x u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t64 q a c x p s f b case2 u i)))))))))))). Time Defined. (* constant 5898 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t66 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1), l_con))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_ec3e21 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz18 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t65 q a c x p s f b case2 u i)))))))))))). Time Defined. (* constant 5899 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t67 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_ore1 (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz24 (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u)) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t66 q a c x p s f b case2 u t))))))))))). Time Defined. (* constant 5900 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_nat)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_mn (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t67 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5901 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t68 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_islessis1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_mn_th1c (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t67 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u)))))))))))). Time Defined. (* constant 5902 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t69 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_or_th9 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) x) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t68 q a c x p s f b case2 u) (fun (t:l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_satz20c (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) x l_e_st_eq_landau_n_1 t) (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_satz20b (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) x l_e_st_eq_landau_n_1 t))))))))))). Time Defined. (* constant 5903 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t69 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5904 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t70 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t69 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5905 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t71 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_mn_th1a (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t67 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm2 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t70 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)) x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)))))))))))))). Time Defined. (* constant 5906 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t72 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t71 q a c x p s f b case2 u)))))))))))). Time Defined. (* constant 5907 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t73 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_thinvf2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s b (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_u4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t72 q a c x p s f b case2 u)))))))))))). Time Defined. (* constant 5908 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t74 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u)) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_1top x u))) (l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2))) (l_e_st_eq_landau_n_isinni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t73 q a c x p s f b case2 u)))))))))))). Time Defined. (* constant 5909 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t75 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_mn_th1e (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t54 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t74 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5910 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t76 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to x) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris (l_e_st_eq_landau_n_1to x) u (l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_1top x u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_isoutinn x u) (l_e_st_eq_landau_n_isoutni x (l_e_st_eq_landau_n_rt_rp_r_c_inn x u) (l_e_st_eq_landau_n_1top x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t55 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t75 q a c x p s f b case2 u)))))))))))). Time Defined. (* constant 5911 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t77 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_image (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2) u)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_somei (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_is (l_e_st_eq_landau_n_1to x) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 t)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_w4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t76 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5912 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t78 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_andi (l_e_injective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2)) (l_e_surjective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2)) (fun (t:l_e_st_eq_landau_n_1to x) => (fun (u:l_e_st_eq_landau_n_1to x) => (fun (v:l_e_is (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2 t) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2 u)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t61 q a c x p s f b case2 t u v))) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t77 q a c x p s f b case2 t)))))))))). Time Defined. (* constant 5913 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) f))))))))). Time Defined. (* constant 5914 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_right l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2)))))))))). Time Defined. (* constant 5915 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t79 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c x (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t78 q a c x p s f b case2)))))))))). Time Defined. (* constant 5916 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))). Time Defined. (* constant 5917 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_right l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2)))))))))). Time Defined. (* constant 5918 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g5 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c x (l_e_st_eq_landau_n_rt_rp_r_c_8283_s03 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2)))))))))). Time Defined. (* constant 5919 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t80 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_mn_th1a (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t54 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_nm1 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t55 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u)) x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u)))))))))))))). Time Defined. (* constant 5920 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t81 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)))) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_n3 q a c x p s f b case2 u) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t80 q a c x p s f b case2 u)))))))))))). Time Defined. (* constant 5921 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t82 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (u:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g5 q a c x p s f b case2 u))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (u:l_e_st_eq_landau_n_1to x) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_rt_rp_r_c_8283_s02 q a c x p s f b case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_u3 q a c x p s f b case2 u)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_w3 q a c x p s f b case2 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t81 q a c x p s f b case2 u))))))))))). Time Defined. (* constant 5922 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t83 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g5 q a c x p s f b case2)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_fisi (l_e_st_eq_landau_n_1to x) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g5 q a c x p s f b case2) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t82 q a c x p s f b case2 t)))))))))). Time Defined. (* constant 5923 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t85 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g5 q a c x p s f b case2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (u:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q x u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g5 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t83 q a c x p s f b case2)))))))))). Time Defined. (* constant 5924 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t86 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g5 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t85 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t79 q a c x p s f b case2)))))))))). Time Defined. (* constant 5925 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t87 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_lessisi1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_satz18a l_e_st_eq_landau_n_1 x)))))))))). Time Defined. (* constant 5926 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t87 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2)))))))))). Time Defined. (* constant 5927 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t87 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2)))))))))). Time Defined. (* constant 5928 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t87a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_isinoutn l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_lessisi3 l_e_st_eq_landau_n_1)))))))))). Time Defined. (* constant 5929 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_1d : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t87 q a c x p s f b case2) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))))))). Time Defined. (* constant 5930 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t87b : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1d q a c x p s f b case2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_1top l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t87 q a c x p s f b case2))))))))))). Time Defined. (* constant 5931 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t88 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1d q a c x p s f b case2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tris l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1d q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t87a q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t87b q a c x p s f b case2)))))))))). Time Defined. (* constant 5932 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_1e : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1d q a c x p s f b case2)))))))))). Time Defined. (* constant 5933 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t88a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1d q a c x p s f b case2)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1d q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1d q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t47 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t88 q a c x p s f b case2)))))))))). Time Defined. (* constant 5934 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t88b : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_symis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t88a q a c x p s f b case2))))))))))). Time Defined. (* constant 5935 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t89 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tr3is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t88b q a c x p s f b case2) case2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t88a q a c x p s f b case2)))))))))). Time Defined. (* constant 5936 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t89a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_satz277 q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1e q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t89 q a c x p s f b case2)) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_satz277 q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2)))))))))))). Time Defined. (* constant 5937 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t90 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_satz281 q a l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2)))))))))). Time Defined. (* constant 5938 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t91 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t86 q a c x p s f b case2)))))))))). Time Defined. (* constant 5939 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t92 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t89a q a c x p s f b case2)))))))))). Time Defined. (* constant 5940 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t93 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2))) (l_e_st_eq_landau_n_rt_rp_r_c_satz281 q a l_e_st_eq_landau_n_1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2))))))))))). Time Defined. (* constant 5941 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t94 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g4 q a c x p s f b case2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g6 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f04 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_f03 q a c x p s f b case2))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t90 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t91 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t92 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t93 q a c x p s f b case2)))))))))). Time Defined. (* constant 5942 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t95 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_compl l_e_st_eq_landau_n_1 x)))))))))). Time Defined. (* constant 5943 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t96 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_issmpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_compl l_e_st_eq_landau_n_1 x))))))))))). Time Defined. (* constant 5944 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t97 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (case2:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g3 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f02 q a c x p s f b case2)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t96 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t94 q a c x p s f b case2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t95 q a c x p s f b case2)))))))))). Time Defined. (* constant 5945 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). Time Defined. (* constant 5946 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_invf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s b (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). Time Defined. (* constant 5947 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t98 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_thinvf2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s b (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). Time Defined. (* constant 5948 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t99 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => not2)))))))))). Time Defined. (* constant 5949 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t100 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_symnotis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (l_e_notis_th3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t98 q a c x p s f b not1 not2)))))))))))). Time Defined. (* constant 5950 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t101 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t100 q a c x p s f b not1 not2) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) t))))))))))). Time Defined. (* constant 5951 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_changef (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))))))))))). Time Defined. (* constant 5952 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t102 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_changef1 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) u i)))))))))))). Time Defined. (* constant 5953 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t103 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t102 q a c x p s f b not1 not2 u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t98 q a c x p s f b not1 not2))))))))))))). Time Defined. (* constant 5954 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t104 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) => l_e_changef2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) u i)))))))))))). Time Defined. (* constant 5955 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t105 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (s u)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_changef3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) u n o))))))))))))). Time Defined. (* constant 5956 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t106 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_wissel_th6 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) b)))))))))). Time Defined. (* constant 5957 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_wissel (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2)))))))))))). Time Defined. (* constant 5958 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t107 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_wissel_th3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2)))))))))))). Time Defined. (* constant 5959 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t108 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_iswissel1 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t103 q a c x p s f b not1 not2 u i)))))))))))))). Time Defined. (* constant 5960 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t109 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t108 q a c x p s f b not1 not2 alpha u i)))))))))))))). Time Defined. (* constant 5961 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t110 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) => l_e_iswissel2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t104 q a c x p s f b not1 not2 u i)))))))))))))). Time Defined. (* constant 5962 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t111 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) => l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t98 q a c x p s f b not1 not2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t110 q a c x p s f b not1 not2 alpha u i)))))))))))))). Time Defined. (* constant 5963 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t112 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_isfe (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t8 q a c x p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f (l_e_st_eq_landau_n_rt_rp_r_c_8283_t106 q a c x p s f b not1 not2)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i (l_e_st_eq_landau_n_rt_rp_r_c_8283_t103 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_refis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))))))). Time Defined. (* constant 5964 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t113 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) n (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t112 q a c x p s f b not1 not2 alpha u n o t))))))))))))))). Time Defined. (* constant 5965 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t114 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2)) => l_e_isfe (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t8 q a c x p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f (l_e_st_eq_landau_n_rt_rp_r_c_8283_t106 q a c x p s f b not1 not2)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) i (l_e_st_eq_landau_n_rt_rp_r_c_8283_t104 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_refis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))))))))))))))))))). Time Defined. (* constant 5966 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t115 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2)) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) o (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t114 q a c x p s f b not1 not2 alpha u n o t))))))))))))))). Time Defined. (* constant 5967 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t116 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_iswissel3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t113 q a c x p s f b not1 not2 alpha u n o) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t115 q a c x p s f b not1 not2 alpha u n o))))))))))))))). Time Defined. (* constant 5968 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t117 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_symis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)) (s u) (l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t116 q a c x p s f b not1 not2 alpha u n o) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t105 q a c x p s f b not1 not2 u n o)))))))))))))))). Time Defined. (* constant 5969 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t118 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u))) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t111 q a c x p s f b not1 not2 alpha u t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t117 q a c x p s f b not1 not2 alpha u n t)))))))))))))). Time Defined. (* constant 5970 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t119 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u))) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t109 q a c x p s f b not1 not2 alpha u t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t118 q a c x p s f b not1 not2 alpha u t))))))))))))). Time Defined. (* constant 5971 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t120 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 t))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_fisi (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 t)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t119 q a c x p s f b not1 not2 alpha t)))))))))))). Time Defined. (* constant 5972 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t121 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t103 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_refis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). Time Defined. (* constant 5973 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t121a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_ec3e21 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ismore1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_1px q a c x p s f b) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_compl l_e_st_eq_landau_n_1 x) (l_e_st_eq_landau_n_satz18 l_e_st_eq_landau_n_1 x)))))))))))). Time Defined. (* constant 5974 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t122 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_imp_th3 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t121a q a c x p s f b not1 not2) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) t))))))))))). Time Defined. (* constant 5975 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t123 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_symnotis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) alpha))))))))))). Time Defined. (* constant 5976 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t124 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_iswissel3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t122 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t123 q a c x p s f b not1 not2 alpha)))))))))))). Time Defined. (* constant 5977 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t125 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha) f)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) u f)) s (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 t)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t120 q a c x p s f b not1 not2 alpha)))))))))))). Time Defined. (* constant 5978 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t126 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha) f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha) f))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t97 q a c x p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha) f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t106 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t121 q a c x p s f b not1 not2)))))))))))). Time Defined. (* constant 5979 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t127 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t46 q a c x p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha) f (l_e_st_eq_landau_n_rt_rp_r_c_8283_t107 q a c x p s f b not1 not2 alpha) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t124 q a c x p s f b not1 not2 alpha)))))))))))). Time Defined. (* constant 5980 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t128 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (alpha:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha) f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s2 q a c x p s f b not1 not2 alpha) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t125 q a c x p s f b not1 not2 alpha) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t126 q a c x p s f b not1 not2 alpha) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t127 q a c x p s f b not1 not2 alpha)))))))))))). Time Defined. (* constant 5981 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_wissel (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))))))))))))). Time Defined. (* constant 5982 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t129 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_wissel_th3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))))))))))))). Time Defined. (* constant 5983 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t130 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t104 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u) (l_e_iswissel1 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) u i))))))))))))))). Time Defined. (* constant 5984 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t131 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t130 q a c x p s f b not1 not2 i3 beta u i))))))))))))))). Time Defined. (* constant 5985 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t132 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t103 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u) (l_e_iswissel2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) u i))))))))))))))). Time Defined. (* constant 5986 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t133 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) => l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) i) (l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t132 q a c x p s f b not1 not2 i3 beta u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t98 q a c x p s f b not1 not2)))))))))))))))). Time Defined. (* constant 5987 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t134 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u) u))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_iswissel3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) u n o))))))))))))))). Time Defined. (* constant 5988 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t135 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_t134 q a c x p s f b not1 not2 i3 beta u n o)))))))))))))))). Time Defined. (* constant 5989 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t136 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (s u)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t105 q a c x p s f b not1 not2 u n o))))))))))))))). Time Defined. (* constant 5990 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t139 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (o:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_symis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)) (s u) (l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 u) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t135 q a c x p s f b not1 not2 i3 beta u n o) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t136 q a c x p s f b not1 not2 i3 beta u n o))))))))))))))))). Time Defined. (* constant 5991 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t140 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), (forall (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => (fun (n:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u))) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t133 q a c x p s f b not1 not2 i3 beta u t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t139 q a c x p s f b not1 not2 i3 beta u n t))))))))))))))). Time Defined. (* constant 5992 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t141 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta u))) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t131 q a c x p s f b not1 not2 i3 beta u t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t140 q a c x p s f b not1 not2 i3 beta u t)))))))))))))). Time Defined. (* constant 5993 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t142 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta t)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_fisi (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta t)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t141 q a c x p s f b not1 not2 i3 beta t))))))))))))). Time Defined. (* constant 5994 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t143 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_symnotis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) beta)))))))))))). Time Defined. (* constant 5995 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t144 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_iswissel3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t122 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t143 q a c x p s f b not1 not2 i3 beta))))))))))))). Time Defined. (* constant 5996 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t145 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) u f)) s (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta t)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t142 q a c x p s f b not1 not2 i3 beta))))))))))))). Time Defined. (* constant 5997 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t146 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t46 q a c x p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t129 q a c x p s f b not1 not2 i3 beta) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t144 q a c x p s f b not1 not2 i3 beta))))))))))))). Time Defined. (* constant 5998 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t147 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t97 q a c x p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f (l_e_st_eq_landau_n_rt_rp_r_c_8283_t106 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t121 q a c x p s f b not1 not2))))))))))))). Time Defined. (* constant 5999 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t148 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (beta:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s3 q a c x p s f b not1 not2 i3 beta) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s1 q a c x p s f b not1 not2) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t145 q a c x p s f b not1 not2 i3 beta) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t146 q a c x p s f b not1 not2 i3 beta) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t147 q a c x p s f b not1 not2 i3 beta))))))))))))). Time Defined. (* constant 6000 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t149 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_ispl1 l_e_st_eq_landau_n_1 x l_e_st_eq_landau_n_1 (l_e_symis l_e_st_eq_landau_n_nat x l_e_st_eq_landau_n_1 i)))))))))))))). Time Defined. (* constant 6001 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_lessisi2 l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t149 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6002 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) f))))))))))))). Time Defined. (* constant 6003 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) s))))))))))))). Time Defined. (* constant 6004 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))))))). Time Defined. (* constant 6005 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t151 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t149 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6006 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t152 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t149 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6007 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t153 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_satz280 q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6008 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t154 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_satz280 q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6009 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t155 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_tris1 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_isinoutn l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_lessisi3 l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t149 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6010 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t156 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t155 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6011 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t157 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t156 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6012 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t158 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) l_e_st_eq_landau_n_1))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_symis l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_isinoutn l_e_st_eq_landau_n_2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a l_e_st_eq_landau_n_2))))))))))))))). Time Defined. (* constant 6013 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t159 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t158 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6014 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t160 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_tr3is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t159 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t98 q a c x p s f b not1 not2) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) gamma)))))))))))))). Time Defined. (* constant 6015 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t161 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t157 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t160 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6016 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t163 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t159 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6017 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t164 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t163 q a c x p s f b not1 not2 i3 gamma i) i3))))))))))))). Time Defined. (* constant 6018 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t165 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t164 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t156 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6019 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t166 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t161 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6020 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t167 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_rt_rp_r_c_8283_s05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t150 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t165 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6021 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t168 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t166 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6022 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t169 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_comq q c (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))))))))))))))). Time Defined. (* constant 6023 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t170 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t167 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6024 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t171 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t152 q a c x p s f b not1 not2 i3 gamma i)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t154 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t168 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t169 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6025 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t172 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g7 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t171 q a c x p s f b not1 not2 i3 gamma i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t170 q a c x p s f b not1 not2 i3 gamma i) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_2 (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_1out l_e_st_eq_landau_n_2)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_f05 q a c x p s f b not1 not2 i3 gamma i (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_2))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t153 q a c x p s f b not1 not2 i3 gamma i)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t151 q a c x p s f b not1 not2 i3 gamma i)))))))))))))). Time Defined. (* constant 6026 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_trivial : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (i:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t172 q a c x p s f b not1 not2 i3 gamma i))))))))))))). Time Defined. (* constant 6027 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t173 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_more x l_e_st_eq_landau_n_1))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_ore1 (l_e_st_eq_landau_n_more x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz24 x) n))))))))))))). Time Defined. (* constant 6028 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_nat))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_mn x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t173 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6029 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_changef (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). Time Defined. (* constant 6030 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t174 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_wissel_th6 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) b))))))))))))). Time Defined. (* constant 6031 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t175 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_changef2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_refis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). Time Defined. (* constant 6032 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t176 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_symis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i3))))))))))))). Time Defined. (* constant 6033 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t177 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tris2 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t175 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t176 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6034 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t178 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))))))))))))). Time Defined. (* constant 6035 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t178 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))))))). Time Defined. (* constant 6036 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t179 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))))))). Time Defined. (* constant 6037 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t180 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_mn_th1b x l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t173 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6038 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t181 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t180 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6039 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g9 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n))), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t181 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6040 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t182 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g9 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g9 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_issmpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t180 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). Time Defined. (* constant 6041 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_right l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g9 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6042 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t183 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_lessisi1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_satz18a l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). Time Defined. (* constant 6043 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t183 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g9 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6044 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t184 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g9 q a c x p s f b not1 not2 i3 gamma n)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_satz281 q a l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g9 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6045 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t185 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6046 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t186 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isinoutn l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_lessisi3 l_e_st_eq_landau_n_1)))))))))))))). Time Defined. (* constant 6047 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_1a : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t183 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))))))))))))). Time Defined. (* constant 6048 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t187 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1a q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_1top l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t183 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). Time Defined. (* constant 6049 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_1b : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to x))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left1to x (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t181 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1a q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6050 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t188 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1a q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_1b q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1a q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1a q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) x (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1a q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t181 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). Time Defined. (* constant 6051 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t189 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_1b q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tr3is l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1a q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_1b q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t186 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t187 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t188 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6052 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_1c : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t178 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1b q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6053 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t190 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1c q a c x p s f b not1 not2 i3 gamma n)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_1b q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_1b q a c x p s f b not1 not2 i3 gamma n)) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8283_1b q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t178 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t189 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6054 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t191 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1c q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t190 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6055 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t192 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t185 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t191 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6056 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t193 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_symis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t175 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6057 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t194 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t193 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6058 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t195 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t192 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t194 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6059 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t196 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t195 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6060 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t197 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g9 q a c x p s f b not1 not2 i3 gamma n)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g11 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t182 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t184 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t196 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6061 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t198 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t197 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6062 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t199 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_assq1 q a (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). Time Defined. (* constant 6063 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t200 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_comq q c (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). Time Defined. (* constant 6064 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t178 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f)))))))))))))). Time Defined. (* constant 6065 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g13 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n))), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t181 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6066 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_right l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g13 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6067 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_lessisi1 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_satz18a l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g13 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6068 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t201 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f)))))))))))))). Time Defined. (* constant 6069 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t202 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g13 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g13 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_issmpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t180 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). Time Defined. (* constant 6070 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t203 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g13 q a c x p s f b not1 not2 i3 gamma n)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_satz281 q a l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g13 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6071 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t204 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6072 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t205 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_1c q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t190 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6073 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t206 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t204 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t205 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6074 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t207 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_changef1 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_refis (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). Time Defined. (* constant 6075 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t208 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t207 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6076 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t209 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t206 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t208 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6077 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t210 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t209 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6078 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t211 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g13 q a c x p s f b not1 not2 i3 gamma n)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_g15 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t202 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t203 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t210 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6079 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_ua : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => l_e_st_eq_landau_n_right1to l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)))))))))))))). Time Defined. (* constant 6080 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_ub : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_1to x)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => l_e_st_eq_landau_n_left1to x (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t181 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_ua q a c x p s f b not1 not2 i3 gamma n u))))))))))))))). Time Defined. (* constant 6081 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_uc : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_t178 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u))))))))))))))). Time Defined. (* constant 6082 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t212 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t178 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i))))))))))))))). Time Defined. (* constant 6083 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t213 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_satz16a (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))))))))))))))). Time Defined. (* constant 6084 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t214 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_con))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_ec3e31 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t213 q a c x p s f b not1 not2 i3 gamma n u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t212 q a c x p s f b not1 not2 i3 gamma n u i)))))))))))))))). Time Defined. (* constant 6085 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t215 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t214 q a c x p s f b not1 not2 i3 gamma n u t))))))))))))))). Time Defined. (* constant 6086 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t216 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) l_e_st_eq_landau_n_1))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_isoutne (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) x (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t178 q a c x p s f b not1 not2 i3 gamma n)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_satz24a (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) i))))))))))))))). Time Defined. (* constant 6087 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t217 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_ua q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_ua q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_ua q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) x (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_ua q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t181 q a c x p s f b not1 not2 i3 gamma n))))))))))))))))). Time Defined. (* constant 6088 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t218 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_ua q a c x p s f b not1 not2 i3 gamma n u))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u))))))))))))))))). Time Defined. (* constant 6089 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t219 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) l_e_st_eq_landau_n_1))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_tr3is l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_ua q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8283_ub q a c x p s f b not1 not2 i3 gamma n u)) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_8283_t218 q a c x p s f b not1 not2 i3 gamma n u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t217 q a c x p s f b not1 not2 i3 gamma n u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t216 q a c x p s f b not1 not2 i3 gamma n u i)))))))))))))))). Time Defined. (* constant 6090 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t220 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) l_e_st_eq_landau_n_1))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_satz18 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)))))))))))))))). Time Defined. (* constant 6091 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t221 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), (forall (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_con))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (i:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_ec3e21 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_more (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz10b (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t220 q a c x p s f b not1 not2 i3 gamma n u i) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t219 q a c x p s f b not1 not2 i3 gamma n u i)))))))))))))))). Time Defined. (* constant 6092 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t222 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t221 q a c x p s f b not1 not2 i3 gamma n u t))))))))))))))). Time Defined. (* constant 6093 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t223 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => l_e_changef3 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t222 q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t215 q a c x p s f b not1 not2 i3 gamma n u))))))))))))))). Time Defined. (* constant 6094 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t224 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), (forall (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n u))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u)) (s (l_e_st_eq_landau_n_rt_rp_r_c_8283_uc q a c x p s f b not1 not2 i3 gamma n u)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t223 q a c x p s f b not1 not2 i3 gamma n u))))))))))))))). Time Defined. (* constant 6095 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t225 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_fisi (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n t) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n t) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t224 q a c x p s f b not1 not2 i3 gamma n t))))))))))))))). Time Defined. (* constant 6096 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t226 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) u) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t225 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6097 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t227 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_comq q c (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))))))))))))). Time Defined. (* constant 6098 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t228 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t226 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6099 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t229 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n)))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t227 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t228 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6100 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t230 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g14 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t229 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t211 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6101 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t231 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t230 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6102 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t232 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t201 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6103 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t233 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f)))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t46 q a c x p (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_rt_rp_r_c_8283_t174 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t177 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6104 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t234 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g8 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))))) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t179 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t198 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t199 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t200 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6105 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t235 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (n:l_e_st_eq_landau_n_nis x l_e_st_eq_landau_n_1) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)) (q (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xm1 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g10 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (l_e_st_eq_landau_n_rt_rp_r_c_8283_g12 q a c x p s f b not1 not2 i3 gamma n)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_g q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_s4 q a c x p s f b not1 not2 i3 gamma n) f)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) f) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t234 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t231 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t232 q a c x p s f b not1 not2 i3 gamma n) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t233 q a c x p s f b not1 not2 i3 gamma n)))))))))))))). Time Defined. (* constant 6106 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t236 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (gamma:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_imp_th1 (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f) (fun (t:l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t172 q a c x p s f b not1 not2 i3 gamma t) (fun (t:l_not (l_e_st_eq_landau_n_is x l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t235 q a c x p s f b not1 not2 i3 gamma t))))))))))))). Time Defined. (* constant 6107 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t237 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (i3:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t236 q a c x p s f b not1 not2 i3 t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_b0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t148 q a c x p s f b not1 not2 i3 t)))))))))))). Time Defined. (* constant 6108 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t238 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), (forall (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => (fun (not2:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t237 q a c x p s f b not1 not2 t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_rt_rp_r_c_8283_a0 q a c x p s f b not1 not2) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t128 q a c x p s f b not1 not2 t))))))))))). Time Defined. (* constant 6109 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t239 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), (forall (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f))))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => (fun (not1:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t97 q a c x p s f b t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_1out (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t238 q a c x p s f b not1 t)))))))))). Time Defined. (* constant 6110 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t240 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), (forall (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) s) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_rt_rp_r_c_8283_prop1 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) s f) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t46 q a c x p s f b t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (s (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t239 q a c x p s f b t))))))))). Time Defined. (* constant 6111 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t241 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (v:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (w:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) u) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t240 q a c x p u v w)))))))). Time Defined. (* constant 6112 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t242 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c (l_e_st_eq_landau_n_suc x)))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c t) (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t241 q a c x p) (l_e_st_eq_landau_n_satz4a x)))))). Time Defined. (* constant 6113 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8283_t243 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x)))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c t) (l_e_st_eq_landau_n_rt_rp_r_c_8283_t6 q a c) (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c t) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t242 q a c t u)) x)))). Time Defined. (* constant 6114 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz283 : (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (x:l_e_st_eq_landau_n_nat), (forall (s:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)), (forall (b:l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) s), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x (fun (t:l_e_st_eq_landau_n_1to x) => f (s t))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q x f)))))))). exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_1to x)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to x) (l_e_st_eq_landau_n_1to x) s) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8283_t243 q a c x s f b))))))). Time Defined. (* constant 6115 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_shiftl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), l_e_st_eq_landau_n_nat))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => l_e_st_eq_landau_n_rt_rp_r_shiftl x ix y iy ly))))). Time Defined. (* constant 6116 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_shiftr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_real)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_shiftr x ix y iy ly n)))))). Time Defined. (* constant 6117 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_intshiftr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_intshiftr x ix y iy ly n)))))). Time Defined. (* constant 6118 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_shiftrls : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n) x)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_shiftrls x ix y iy ly n)))))). Time Defined. (* constant 6119 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_lsshiftr x ix y iy ly n)))))). Time Defined. (* constant 6120 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iseshiftr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), (forall (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly m)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) n m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => (fun (m:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly m)) => l_e_st_eq_landau_n_rt_rp_r_iseshiftr x ix y iy ly n m i)))))))). Time Defined. (* constant 6121 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_shiftl1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_rt_rp_r_shiftl1 x ix y iy ly u a))))))). Time Defined. (* constant 6122 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_shiftinv1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_is u (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_shiftl1 x ix y iy ly u a))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_rt_rp_r_shiftinv1 x ix y iy ly u a))))))). Time Defined. (* constant 6123 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_shiftinv2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_shiftl1 x ix y iy ly u a)) u))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (a:l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl u) (l_e_st_eq_landau_n_rt_rp_r_lessis y u) (l_e_st_eq_landau_n_rt_rp_r_lessis u x)) => l_e_st_eq_landau_n_rt_rp_r_shiftinv2 x ix y iy ly u a))))))). Time Defined. (* constant 6124 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_shiftf : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_shiftf x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f)))))). Time Defined. (* constant 6125 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_smpri : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f)))))))). Time Defined. (* constant 6126 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (lv:l_e_st_eq_landau_n_rt_rp_r_lessis y v), (forall (kv:l_e_st_eq_landau_n_rt_rp_r_lessis v u), l_e_st_eq_landau_n_rt_rp_r_lessis v x))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (lv:l_e_st_eq_landau_n_rt_rp_r_lessis y v) => (fun (kv:l_e_st_eq_landau_n_rt_rp_r_lessis v u) => l_e_st_eq_landau_n_rt_rp_r_lessisi1 v x (l_e_st_eq_landau_n_rt_rp_r_satz172a v u x kv k)))))))))))))))))). Time Defined. (* constant 6127 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lft : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (v:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (lt:l_e_st_eq_landau_n_rt_rp_r_lessis y t), (forall (kt:l_e_st_eq_landau_n_rt_rp_r_lessis t u), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_intrl t) => (fun (lt:l_e_st_eq_landau_n_rt_rp_r_lessis y t) => (fun (kt:l_e_st_eq_landau_n_rt_rp_r_lessis t u) => f t v lt (l_e_st_eq_landau_n_rt_rp_r_c_8284_t1 x ix y iy ly f pi q a u iu l k t v lt kt)))))))))))))))))). Time Defined. (* constant 6128 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_pl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl x y)). Time Defined. (* constant 6129 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_mn : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_mn x y)). Time Defined. (* constant 6130 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl). Time Defined. (* constant 6131 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (lv:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) v), (forall (kv:l_e_st_eq_landau_n_rt_rp_r_lessis v x), l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl u l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (lv:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) v) => (fun (kv:l_e_st_eq_landau_n_rt_rp_r_lessis v x) => l_e_st_eq_landau_n_rt_rp_r_satz190c u u l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_lessisi2 u u (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real u)) (l_e_st_eq_landau_n_rt_rp_r_lemma1 l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satz169a l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_natpos l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_natrl1)))))))))))))))))))). Time Defined. (* constant 6132 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (lv:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) v), (forall (kv:l_e_st_eq_landau_n_rt_rp_r_lessis v x), l_e_st_eq_landau_n_rt_rp_r_less u (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (lv:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) v) => (fun (kv:l_e_st_eq_landau_n_rt_rp_r_lessis v x) => l_e_st_eq_landau_n_rt_rp_r_isless1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl u l_e_st_eq_landau_n_rt_rp_r_0) u (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_pl02 u l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t2 x ix y iy ly f pi q a u iu l k v iv lv kv)))))))))))))))))). Time Defined. (* constant 6133 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (lv:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) v), (forall (kv:l_e_st_eq_landau_n_rt_rp_r_lessis v x), l_e_st_eq_landau_n_rt_rp_r_lessis y v))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (lv:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) v) => (fun (kv:l_e_st_eq_landau_n_rt_rp_r_lessis v x) => l_e_st_eq_landau_n_rt_rp_r_lessisi1 y v (l_e_st_eq_landau_n_rt_rp_r_satz172b y (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) v (l_e_st_eq_landau_n_rt_rp_r_satz172a y u (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) l (l_e_st_eq_landau_n_rt_rp_r_c_8284_t3 x ix y iy ly f pi q a u iu l k v iv lv kv)) lv)))))))))))))))))). Time Defined. (* constant 6134 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_rgt : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (v:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (lt:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) t), (forall (kt:l_e_st_eq_landau_n_rt_rp_r_lessis t x), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_intrl t) => (fun (lt:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) t) => (fun (kt:l_e_st_eq_landau_n_rt_rp_r_lessis t x) => f t v (l_e_st_eq_landau_n_rt_rp_r_c_8284_t4 x ix y iy ly f pi q a u iu l k t v lt kt) kt))))))))))))))))). Time Defined. (* constant 6135 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_intpl u iu l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1))))))))))))). Time Defined. (* constant 6136 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) x))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_satzr25 u iu x ix k))))))))))))). Time Defined. (* constant 6137 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) y)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) y) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u))) (l_e_st_eq_landau_n_rt_rp_r_asspl2 (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) y (l_e_st_eq_landau_n_rt_rp_r_plmn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)))))))))))))))). Time Defined. (* constant 6138 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_nat))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly))))))))))))). Time Defined. (* constant 6139 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_suy : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_nat))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_c_shiftl u iu y iy l))))))))))))). Time Defined. (* constant 6140 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_nat))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6141 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_satzr155a (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_isrlnt2 (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 u iu y iy l)) (l_e_st_eq_landau_n_rt_rp_r_isrlnt2 (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k)))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t7 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 x) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 x ix y iy ly))))))))))))))). Time Defined. (* constant 6142 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_isntirl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t8 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6143 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t9 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6144 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k))), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t10 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f)))))))))))))). Time Defined. (* constant 6145 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t9 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6146 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_fr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_right l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6147 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_lessisi1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_satz18a (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k))))))))))))))). Time Defined. (* constant 6148 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_fl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6149 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t12a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 x ix y iy ly f pi q a u iu l k)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_st_eq_landau_n_rt_rp_r_c_satz281 q a (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6150 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_right1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_satz19o (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)))))))))))))))). Time Defined. (* constant 6151 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_right1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t10 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_right1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_right1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_right1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_right1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t10 x ix y iy ly f pi q a u iu l k)))))))))))))))). Time Defined. (* constant 6152 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t10 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_right1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))))))))))))))). Time Defined. (* constant 6153 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t15 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_right1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t13 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t14 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6154 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t16 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_isnterl (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t15 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6155 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t17 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_satzr155b (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))))))))))))))). Time Defined. (* constant 6156 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t18 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 u iu y iy l)))))))))))))))). Time Defined. (* constant 6157 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t19 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t18 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t17 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t16 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6158 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t20 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))) y))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))) y (l_e_st_eq_landau_n_rt_rp_r_c_8284_t19 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6159 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t21 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y)) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) y)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y)) y (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)))) (l_e_st_eq_landau_n_rt_rp_r_asspl1 (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) y) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_plmn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y)))))))))))))))). Time Defined. (* constant 6160 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t22 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_mn (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) y) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n))) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t20 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t21 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6161 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t23 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n)) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u)) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8284_t22 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6162 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t24 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6163 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t25 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6164 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t26 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n)) x)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6165 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t27 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n)))))))))))))). Time Defined. (* constant 6166 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t28 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n)))))))))))))). Time Defined. (* constant 6167 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t29 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n) x)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n)))))))))))))). Time Defined. (* constant 6168 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t30 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_8284_t4 x ix y iy ly f pi q a u iu l k (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t27 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t28 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t29 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6169 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t31 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => pi (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n1 x ix y iy ly f pi q a u iu l k n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t24 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t25 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t26 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t27 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t30 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t29 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t23 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6170 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t32 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_fisi (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_8284_t31 x ix y iy ly f pi q a u iu l k t)))))))))))))). Time Defined. (* constant 6171 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t33 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (v:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) v) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t32 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6172 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t34 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k))) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t33 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6173 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t35 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k)))))))))))))))). Time Defined. (* constant 6174 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t36 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t10 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t10 x ix y iy ly f pi q a u iu l k)))))))))))))))). Time Defined. (* constant 6175 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t10 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) n))))))))))))))). Time Defined. (* constant 6176 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t37 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_tris l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12 x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t35 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t36 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6177 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t38 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_isnterl (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t37 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6178 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t39 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n)) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n))) y))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n))) y (l_e_st_eq_landau_n_rt_rp_r_c_8284_t38 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6179 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t40 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_shiftr u iu y iy l n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) n)) y) (l_e_st_eq_landau_n_rt_rp_r_c_8284_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n))) y) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8284_t39 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6180 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t41 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_shiftr u iu y iy l n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_intshiftr u iu y iy l n)))))))))))))). Time Defined. (* constant 6181 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t42 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_shiftr u iu y iy l n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr u iu y iy l n)))))))))))))). Time Defined. (* constant 6182 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t43 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_shiftr u iu y iy l n) u)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftrls u iu y iy l n)))))))))))))). Time Defined. (* constant 6183 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t44 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_shiftr u iu y iy l n) x)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_8284_t1 x ix y iy ly f pi q a u iu l k (l_e_st_eq_landau_n_rt_rp_r_c_shiftr u iu y iy l n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t41 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t42 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t43 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6184 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t45 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6185 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t46 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6186 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t47 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n)) x)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6187 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t48 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_shiftf u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => pi (l_e_st_eq_landau_n_rt_rp_r_c_shiftr u iu y iy l n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t41 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t42 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t44 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8284_n2 x ix y iy ly f pi q a u iu l k n)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t45 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t46 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t47 x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t40 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6188 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t49 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) n))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_shiftf u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k n) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t48 x ix y iy ly f pi q a u iu l k n))))))))))))))). Time Defined. (* constant 6189 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t50 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_fisi (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)) => l_e_st_eq_landau_n_rt_rp_r_c_8284_t49 x ix y iy ly f pi q a u iu l k t)))))))))))))). Time Defined. (* constant 6190 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t51 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) q)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (v:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) v) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t50 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6191 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t52 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_c_is (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpri u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_isf l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => q t (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) q) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t51 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6192 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8284_t53 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 x ix y iy ly f pi q a u iu l k)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpri u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 x ix y iy ly f pi q a u iu l k)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fr x ix y iy ly f pi q a u iu l k))) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_fl x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q)) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpri u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_c_8284_p1 u) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t5 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t6 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t12a x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t34 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t52 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6193 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz284 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (k:l_e_st_eq_landau_n_rt_rp_r_less u x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpri u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_intpl u iu l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1) (l_e_st_eq_landau_n_rt_rp_r_satzr25 u iu x ix k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (l:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (k:l_e_st_eq_landau_n_rt_rp_r_less u x) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q) (q (l_e_st_eq_landau_n_rt_rp_r_c_smpri u iu y iy l (l_e_st_eq_landau_n_rt_rp_r_c_lft x ix y iy ly f pi q a u iu l k) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix (l_e_st_eq_landau_n_rt_rp_r_pl u l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_intpl u iu l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1) (l_e_st_eq_landau_n_rt_rp_r_satzr25 u iu x ix k) (l_e_st_eq_landau_n_rt_rp_r_c_rgt x ix y iy ly f pi q a u iu l k) q)) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_8284_suy x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_sxu x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_f1 x ix y iy ly f pi q a u iu l k)) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t11 x ix y iy ly f pi q a u iu l k) (l_e_st_eq_landau_n_rt_rp_r_c_8284_t53 x ix y iy ly f pi q a u iu l k)))))))))))))). Time Defined. (* constant 6194 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_pl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl x y)). Time Defined. (* constant 6195 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_mn : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_mn x y)). Time Defined. (* constant 6196 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl x l_e_st_eq_landau_n_rt_rp_r_1rl). Time Defined. (* constant 6197 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (w:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iw:l_e_st_eq_landau_n_rt_rp_r_intrl w), (forall (lw:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) w), (forall (kw:l_e_st_eq_landau_n_rt_rp_r_lessis w (l_e_st_eq_landau_n_rt_rp_r_pl x v)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iw:l_e_st_eq_landau_n_rt_rp_r_intrl w) => (fun (lw:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) w) => (fun (kw:l_e_st_eq_landau_n_rt_rp_r_lessis w (l_e_st_eq_landau_n_rt_rp_r_pl x v)) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) w) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) w) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v)) lw (fun (t:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) w) => l_e_st_eq_landau_n_rt_rp_r_satz188f (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) w (l_e_st_eq_landau_n_rt_rp_r_m0 v) t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) w) => l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) w v t))))))))))))))). Time Defined. (* constant 6198 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (w:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iw:l_e_st_eq_landau_n_rt_rp_r_intrl w), (forall (lw:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) w), (forall (kw:l_e_st_eq_landau_n_rt_rp_r_lessis w (l_e_st_eq_landau_n_rt_rp_r_pl x v)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iw:l_e_st_eq_landau_n_rt_rp_r_intrl w) => (fun (lw:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) w) => (fun (kw:l_e_st_eq_landau_n_rt_rp_r_lessis w (l_e_st_eq_landau_n_rt_rp_r_pl x v)) => l_e_st_eq_landau_n_rt_rp_r_islessis1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) v) y (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v) (l_e_st_eq_landau_n_rt_rp_r_mnpl y v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t1 x ix y iy ly f pi q v iv w iw lw kw))))))))))))))). Time Defined. (* constant 6199 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (w:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iw:l_e_st_eq_landau_n_rt_rp_r_intrl w), (forall (lw:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) w), (forall (kw:l_e_st_eq_landau_n_rt_rp_r_lessis w (l_e_st_eq_landau_n_rt_rp_r_pl x v)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) v))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iw:l_e_st_eq_landau_n_rt_rp_r_intrl w) => (fun (lw:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) w) => (fun (kw:l_e_st_eq_landau_n_rt_rp_r_lessis w (l_e_st_eq_landau_n_rt_rp_r_pl x v)) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_less w (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_is w (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) v)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) v)) kw (fun (t:l_e_st_eq_landau_n_rt_rp_r_less w (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) => l_e_st_eq_landau_n_rt_rp_r_satz188f w (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_m0 v) t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is w (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) => l_e_st_eq_landau_n_rt_rp_r_ismn1 w (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) v t))))))))))))))). Time Defined. (* constant 6200 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (w:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iw:l_e_st_eq_landau_n_rt_rp_r_intrl w), (forall (lw:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) w), (forall (kw:l_e_st_eq_landau_n_rt_rp_r_lessis w (l_e_st_eq_landau_n_rt_rp_r_pl x v)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v) x)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iw:l_e_st_eq_landau_n_rt_rp_r_intrl w) => (fun (lw:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) w) => (fun (kw:l_e_st_eq_landau_n_rt_rp_r_lessis w (l_e_st_eq_landau_n_rt_rp_r_pl x v)) => l_e_st_eq_landau_n_rt_rp_r_islessis2 (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) v) x (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn w v) (l_e_st_eq_landau_n_rt_rp_r_mnpl x v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t3 x ix y iy ly f pi q v iv w iw lw kw))))))))))))))). Time Defined. (* constant 6201 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_sft : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (w:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (lt:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) t), (forall (kt:l_e_st_eq_landau_n_rt_rp_r_lessis t (l_e_st_eq_landau_n_rt_rp_r_pl x v)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_intrl t) => (fun (lt:l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) t) => (fun (kt:l_e_st_eq_landau_n_rt_rp_r_lessis t (l_e_st_eq_landau_n_rt_rp_r_pl x v)) => f (l_e_st_eq_landau_n_rt_rp_r_mn t v) (l_e_st_eq_landau_n_rt_rp_r_intmn t w v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t2 x ix y iy ly f pi q v iv t w lt kt) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t4 x ix y iy ly f pi q v iv t w lt kt))))))))))))))). Time Defined. (* constant 6202 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_m0 v) (l_e_st_eq_landau_n_rt_rp_r_m0 y)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl v y)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_m0 v) (l_e_st_eq_landau_n_rt_rp_r_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl v y) (l_e_st_eq_landau_n_rt_rp_r_compl y v)) (l_e_st_eq_landau_n_rt_rp_r_satz180 v y))))))))))). Time Defined. (* constant 6203 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 x))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl x) (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 x) (l_e_st_eq_landau_n_rt_rp_r_asspl1 l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_m0 v)) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) v) x l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_mnpl x v)) (l_e_st_eq_landau_n_rt_rp_r_compl l_e_st_eq_landau_n_rt_rp_r_1rl x))))))))))). Time Defined. (* constant 6204 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 x) y))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_m0 v) (l_e_st_eq_landau_n_rt_rp_r_m0 y))) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) v) y) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 x) y) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_m0 v) (l_e_st_eq_landau_n_rt_rp_r_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t5 x ix y iy ly f pi q v iv)) (l_e_st_eq_landau_n_rt_rp_r_asspl2 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_m0 v) (l_e_st_eq_landau_n_rt_rp_r_m0 y)) (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 x) y (l_e_st_eq_landau_n_rt_rp_r_c_8285_t6 x ix y iy ly f pi q v iv)))))))))))). Time Defined. (* constant 6205 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_less y x) (l_e_st_eq_landau_n_rt_rp_r_is y x) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) ly (fun (t:l_e_st_eq_landau_n_rt_rp_r_less y x) => l_e_st_eq_landau_n_rt_rp_r_satz188f y x v t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is y x) => l_e_st_eq_landau_n_rt_rp_r_ispl1 y x v t))))))))))). Time Defined. (* constant 6206 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_nat)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)))))))))). Time Defined. (* constant 6207 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_sv : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_nat)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_st_eq_landau_n_rt_rp_r_c_shiftl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv))))))))))). Time Defined. (* constant 6208 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_st_eq_landau_n_rt_rp_r_isrlent (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_p1 x) y) (l_e_st_eq_landau_n_rt_rp_r_shift_t6 x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t7 x ix y iy ly f pi q v iv))))))))))). Time Defined. (* constant 6209 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t9 x ix y iy ly f pi q v iv))))))))))). Time Defined. (* constant 6210 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_f1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_c_cx))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t10 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f))))))))))). Time Defined. (* constant 6211 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_f1 x ix y iy ly f pi q v iv)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr q (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f) (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t9 x ix y iy ly f pi q v iv))))))))))). Time Defined. (* constant 6212 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t10 x ix y iy ly f pi q v iv) n))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_isinoutn (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t10 x ix y iy ly f pi q v iv))))))))))))). Time Defined. (* constant 6213 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t10 x ix y iy ly f pi q v iv) n))))))))))). Time Defined. (* constant 6214 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_isnterl (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t12 x ix y iy ly f pi q v iv n)))))))))))). Time Defined. (* constant 6215 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n))) y)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n))) y) (l_e_st_eq_landau_n_rt_rp_r_asspl1 (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_m0 v)) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n))) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) v) y (l_e_st_eq_landau_n_rt_rp_r_c_8285_t13 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_mnpl y v))))))))))))). Time Defined. (* constant 6216 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_st0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_real))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n)))))))))))). Time Defined. (* constant 6217 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_stv : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_real))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftr (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) n))))))))))). Time Defined. (* constant 6218 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t15 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_st0 x ix y iy ly f pi q v iv n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 v))) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_m0 v) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl))) (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) v) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_8285_st0 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_asspl1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 v)) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 v)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_m0 v) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_compl (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_m0 v))) (l_e_st_eq_landau_n_rt_rp_r_asspl2 (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) (l_e_st_eq_landau_n_rt_rp_r_m0 v) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ismn1 (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) n)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v)) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_rt_rp_r_c_8285_s0 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n))) y) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_8285_t14 x ix y iy ly f pi q v iv n))))))))))))). Time Defined. (* constant 6219 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t16 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_8285_st0 x ix y iy ly f pi q v iv n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n)))))))))))). Time Defined. (* constant 6220 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t17 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_st0 x ix y iy ly f pi q v iv n) x))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n)))))))))))). Time Defined. (* constant 6221 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t18 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_8285_st0 x ix y iy ly f pi q v iv n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8285_n1 x ix y iy ly f pi q v iv n)))))))))))). Time Defined. (* constant 6222 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t19 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_intshiftr (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) n))))))))))). Time Defined. (* constant 6223 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t20 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftrls (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) n))))))))))). Time Defined. (* constant 6224 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t21 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) n))))))))))). Time Defined. (* constant 6225 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t22 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) v)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_intmn (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t19 x ix y iy ly f pi q v iv n) v iv))))))))))). Time Defined. (* constant 6226 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t23 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) v)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_8285_t2 x ix y iy ly f pi q v iv (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t19 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t21 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t20 x ix y iy ly f pi q v iv n)))))))))))). Time Defined. (* constant 6227 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t24 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) v) x))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_8285_t4 x ix y iy ly f pi q v iv (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t19 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t21 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t20 x ix y iy ly f pi q v iv n)))))))))))). Time Defined. (* constant 6228 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t25 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_shiftf (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_sft x ix y iy ly f pi q v iv) n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_f1 x ix y iy ly f pi q v iv n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => pi (l_e_st_eq_landau_n_rt_rp_r_c_8285_mn (l_e_st_eq_landau_n_rt_rp_r_c_8285_stv x ix y iy ly f pi q v iv n) v) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t22 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t23 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t24 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_st0 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t16 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t18 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t17 x ix y iy ly f pi q v iv n) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t15 x ix y iy ly f pi q v iv n)))))))))))). Time Defined. (* constant 6229 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t26 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_sft x ix y iy ly f pi q v iv)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_f1 x ix y iy ly f pi q v iv))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_fisi (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_shiftf (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_sft x ix y iy ly f pi q v iv)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_f1 x ix y iy ly f pi q v iv) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)) => l_e_st_eq_landau_n_rt_rp_r_c_8285_t25 x ix y iy ly f pi q v iv t))))))))))). Time Defined. (* constant 6230 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8285_t27 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpri (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_sft x ix y iy ly f pi q v iv) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_f1 x ix y iy ly f pi q v iv)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (w:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) w) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_sft x ix y iy ly f pi q v iv)) (l_e_st_eq_landau_n_rt_rp_r_c_8285_f1 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t26 x ix y iy ly f pi q v iv))))))))))). Time Defined. (* constant 6231 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma285 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl y v) (l_e_st_eq_landau_n_rt_rp_r_pl x v))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_st_eq_landau_n_rt_rp_r_c_8285_t8 x ix y iy ly f pi q v iv)))))))))). Time Defined. (* constant 6232 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz285 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (v:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpri (l_e_st_eq_landau_n_rt_rp_r_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_lemma285 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_sft x ix y iy ly f pi q v iv) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iv:l_e_st_eq_landau_n_rt_rp_r_intrl v) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpri (l_e_st_eq_landau_n_rt_rp_r_pl x v) (l_e_st_eq_landau_n_rt_rp_r_intpl x ix v iv) (l_e_st_eq_landau_n_rt_rp_r_pl y v) (l_e_st_eq_landau_n_rt_rp_r_intpl y iy v iv) (l_e_st_eq_landau_n_rt_rp_r_c_lemma285 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_sft x ix y iy ly f pi q v iv) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_8285_sv x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_f1 x ix y iy ly f pi q v iv)) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t27 x ix y iy ly f pi q v iv) (l_e_st_eq_landau_n_rt_rp_r_c_8285_t11 x ix y iy ly f pi q v iv))))))))))). Time Defined. (* constant 6233 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_us : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x), l_e_st_eq_landau_n_rt_rp_r_real)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x) => s u iu lu ul)))))))))))). Time Defined. (* constant 6234 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x), l_and3 (l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul)) (l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul)) (l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul) x))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x) => ins u iu lu ul)))))))))))). Time Defined. (* constant 6235 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_inseqe1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x) => l_e_st_eq_landau_n_rt_rp_r_shift_t22 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t1 x ix y iy ly s ins f u iu lu ul))))))))))))). Time Defined. (* constant 6236 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_inseqe2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x) => l_e_st_eq_landau_n_rt_rp_r_shift_t23 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t1 x ix y iy ly s ins f u iu lu ul))))))))))))). Time Defined. (* constant 6237 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_inseqe3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul) x)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x) => l_e_st_eq_landau_n_rt_rp_r_shift_t24 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t1 x ix y iy ly s ins f u iu lu ul))))))))))))). Time Defined. (* constant 6238 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_usf : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u), (forall (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u), (forall (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iu:l_e_st_eq_landau_n_rt_rp_r_intrl u) => (fun (lu:l_e_st_eq_landau_n_rt_rp_r_lessis y u) => (fun (ul:l_e_st_eq_landau_n_rt_rp_r_lessis u x) => f (l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f u iu lu ul) (l_e_st_eq_landau_n_rt_rp_r_c_inseqe1 x ix y iy ly s ins f u iu lu ul) (l_e_st_eq_landau_n_rt_rp_r_c_inseqe2 x ix y iy ly s ins f u iu lu ul) (l_e_st_eq_landau_n_rt_rp_r_c_inseqe3 x ix y iy ly s ins f u iu lu ul))))))))))))). Time Defined. (* constant 6239 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_permseq : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_intrl t), (forall (v:l_e_st_eq_landau_n_rt_rp_r_lessis y t), (forall (w:l_e_st_eq_landau_n_rt_rp_r_lessis t x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_intrl t) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_lessis y t) => (fun (w:l_e_st_eq_landau_n_rt_rp_r_lessis t x) => l_e_st_eq_landau_n_rt_rp_r_c_usf x ix y iy ly s ins f t u v w)))))))))))). Time Defined. (* constant 6240 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_ss : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => l_e_st_eq_landau_n_rt_rp_r_shiftseq x ix y iy ly s ins)))))))))))))). Time Defined. (* constant 6241 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps t))) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => l_e_st_eq_landau_n_rt_rp_r_c_satz283 q a c (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly) (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps) (l_e_st_eq_landau_n_rt_rp_r_bijshiftseq x ix y iy ly s ins pri ps) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f))))))))))))))). Time Defined. (* constant 6242 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_ns : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_real))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_us x ix y iy ly s ins f (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix y iy ly n)))))))))))))))). Time Defined. (* constant 6243 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_8286_ns x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftinv1 x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ns x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_shift_t34 x ix y iy ly s ins n)))))))))))))))). Time Defined. (* constant 6244 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_8286_ns x ix y iy ly f pi q a c s ins pri ps n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_inseqe1 x ix y iy ly s ins f (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix y iy ly n)))))))))))))))). Time Defined. (* constant 6245 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_8286_ns x ix y iy ly f pi q a c s ins pri ps n)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_inseqe2 x ix y iy ly s ins f (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix y iy ly n)))))))))))))))). Time Defined. (* constant 6246 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8286_ns x ix y iy ly f pi q a c s ins pri ps n) x))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_inseqe3 x ix y iy ly s ins f (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix y iy ly n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix y iy ly n)))))))))))))))). Time Defined. (* constant 6247 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_intshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n)))))))))))))))). Time Defined. (* constant 6248 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_lessis y (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_lsshiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n)))))))))))))))). Time Defined. (* constant 6249 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n)) x))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftrls x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n)))))))))))))))). Time Defined. (* constant 6250 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_permseq x ix y iy ly s ins f) n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => pi (l_e_st_eq_landau_n_rt_rp_r_c_8286_ns x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t4 x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t5 x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t6 x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_c_shiftr x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps n)) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t7 x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t8 x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t9 x ix y iy ly f pi q a c s ins pri ps n) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t3 x ix y iy ly f pi q a c s ins pri ps n)))))))))))))))). Time Defined. (* constant 6251 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_permseq x ix y iy ly s ins f)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps t)))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => l_e_fisi (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_permseq x ix y iy ly s ins f)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps t)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_8286_t10 x ix y iy ly f pi q a c s ins pri ps t))))))))))))))). Time Defined. (* constant 6252 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8286_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_permseq x ix y iy ly s ins f) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps t))))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly) u) (l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_permseq x ix y iy ly s ins f)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps t)) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t11 x ix y iy ly f pi q a c s ins pri ps))))))))))))))). Time Defined. (* constant 6253 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz286 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_real), (forall (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_real), (forall (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y), (forall (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x), (forall (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f), (forall (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))), (forall (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q), (forall (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q), (forall (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real), (forall (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s), (forall (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s), (forall (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_permseq x ix y iy ly s ins f) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ix:l_e_st_eq_landau_n_rt_rp_r_intrl x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (iy:l_e_st_eq_landau_n_rt_rp_r_intrl y) => (fun (ly:l_e_st_eq_landau_n_rt_rp_r_lessis y x) => (fun (f:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (pi:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_c_cx f) => (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_seq x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real) => (fun (ins:l_e_st_eq_landau_n_rt_rp_r_inseq x ix y iy ly s) => (fun (pri:l_e_st_eq_landau_n_rt_rp_r_proofsirrelevant x ix y iy ly l_e_st_eq_landau_n_rt_rp_r_real s) => (fun (ps:l_e_st_eq_landau_n_rt_rp_r_perm x ix y iy ly s) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly (l_e_st_eq_landau_n_rt_rp_r_c_permseq x ix y iy ly s ins f) q) (l_e_st_eq_landau_n_rt_rp_r_c_smpr q (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_shiftl x ix y iy ly)) => l_e_st_eq_landau_n_rt_rp_r_c_shiftf x ix y iy ly f (l_e_st_eq_landau_n_rt_rp_r_c_8286_ss x ix y iy ly f pi q a c s ins pri ps t))) (l_e_st_eq_landau_n_rt_rp_r_c_smpri x ix y iy ly f q) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t12 x ix y iy ly f pi q a c s ins pri ps) (l_e_st_eq_landau_n_rt_rp_r_c_8286_t2 x ix y iy ly f pi q a c s ins pri ps))))))))))))))). Time Defined. (* constant 6254 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_modf : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (f t)) l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 6255 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x f)) r))). Time Defined. (* constant 6256 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_modf x f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 6257 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 x f r) (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 x f r)))). Time Defined. (* constant 6258 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_prop4 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_some (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x f t))). Time Defined. (* constant 6259 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 : (forall (x:l_e_st_eq_landau_n_nat), Prop). exact (fun (x:l_e_st_eq_landau_n_nat) => (forall (u:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop4 x u)). Time Defined. (* constant 6260 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t1 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum l_e_st_eq_landau_n_1 f) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl t u)) f). Time Defined. (* constant 6261 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t2 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum l_e_st_eq_landau_n_1 f)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_ismod (l_e_st_eq_landau_n_rt_rp_r_c_sum l_e_st_eq_landau_n_1 f) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t1 f)). Time Defined. (* constant 6262 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t3 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 l_e_st_eq_landau_n_1 f (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_lessisi2 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum l_e_st_eq_landau_n_1 f)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t2 f)). Time Defined. (* constant 6263 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t4 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 l_e_st_eq_landau_n_1 f (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl t u)) (l_e_st_eq_landau_n_rt_rp_r_c_modf l_e_st_eq_landau_n_1 f)). Time Defined. (* constant 6264 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t5 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 l_e_st_eq_landau_n_1 f (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 l_e_st_eq_landau_n_1 f (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 l_e_st_eq_landau_n_1 f (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t3 f) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t4 f)). Time Defined. (* constant 6265 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t6 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop4 l_e_st_eq_landau_n_1 f). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 l_e_st_eq_landau_n_1 f t) (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t5 f)). Time Defined. (* constant 6266 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t7 : l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 l_e_st_eq_landau_n_1. exact (fun (u:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8287_t6 u). Time Defined. (* constant 6267 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t8 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))). Time Defined. (* constant 6268 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_lf : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8287_t8 x p f) f))). Time Defined. (* constant 6269 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t9 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) pr))))). Time Defined. (* constant 6270 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t10 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) pr))))). Time Defined. (* constant 6271 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t11 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl t u)) x f))))). Time Defined. (* constant 6272 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_c_ismod (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t11 x p f r pr))))))). Time Defined. (* constant 6273 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_m : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_real))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))))). Time Defined. (* constant 6274 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t13 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_islessis1 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t12 x p f r pr) (l_e_st_eq_landau_n_rt_rp_r_c_satz271 (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))))))). Time Defined. (* constant 6275 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t14 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) r) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) r) (l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))) (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t9 x p f r pr) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) r) => l_e_st_eq_landau_n_rt_rp_r_satz188f (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) r) => l_e_st_eq_landau_n_rt_rp_r_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) t)))))). Time Defined. (* constant 6276 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t15 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t13 x p f r pr) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t14 x p f r pr)))))). Time Defined. (* constant 6277 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_lmf : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8287_t8 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)))))). Time Defined. (* constant 6278 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t16 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lmf x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) l_e_st_eq_landau_n_rt_rp_r_0))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_pl t u)) x (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)))))). Time Defined. (* constant 6279 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t17 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lmf x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) l_e_st_eq_landau_n_rt_rp_r_0))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_c_ispl1 (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lmf x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t10 x p f r pr)))))). Time Defined. (* constant 6280 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t18 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_c_plis12a r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 6281 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t19 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) l_e_st_eq_landau_n_rt_rp_r_0)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))))). Time Defined. (* constant 6282 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t20 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_sum (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lmf x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t16 x p f r pr) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t17 x p f r pr) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t18 x p f r pr) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t19 x p f r pr)))))). Time Defined. (* constant 6283 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t21 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_andi (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr))) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t15 x p f r pr) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t20 x p f r pr)))))). Time Defined. (* constant 6284 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t22 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop4 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (pr:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) r) => l_somei l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f t) (l_e_st_eq_landau_n_rt_rp_r_pl r (l_e_st_eq_landau_n_rt_rp_r_c_8287_m x p f r pr)) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t21 x p f r pr)))))). Time Defined. (* constant 6285 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t23 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop4 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_someapp l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) t) (p (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8287_prop4 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop3 x (l_e_st_eq_landau_n_rt_rp_r_c_8287_lf x p f) t) => l_e_st_eq_landau_n_rt_rp_r_c_8287_t22 x p f t u))))). Time Defined. (* constant 6286 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t25 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8287_t23 x p u))). Time Defined. (* constant 6287 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_t26 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x), l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 (l_e_st_eq_landau_n_suc x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 x) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 t) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_rt_rp_r_c_8287_t25 x p) (l_e_st_eq_landau_n_satz4a x))). Time Defined. (* constant 6288 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz287 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_some (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_and (l_e_st_eq_landau_n_rt_rp_r_lessis (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_sum x f)) t) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_sum x (l_e_st_eq_landau_n_rt_rp_r_c_modf x f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli t l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 t) l_e_st_eq_landau_n_rt_rp_r_c_8287_t7 (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 t) => l_e_st_eq_landau_n_rt_rp_r_c_8287_t26 t u)) x f)). Time Defined. (* constant 6289 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_modf x f)))). Time Defined. (* constant 6290 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 : (forall (x:l_e_st_eq_landau_n_nat), Prop). exact (fun (x:l_e_st_eq_landau_n_nat) => (forall (u:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8288_prop1 x u)). Time Defined. (* constant 6291 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t1 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) f). Time Defined. (* constant 6292 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t2 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_ismod (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t1 f)). Time Defined. (* constant 6293 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t3 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx1 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_t2 f)). Time Defined. (* constant 6294 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t4 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_modf l_e_st_eq_landau_n_1 f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) (l_e_st_eq_landau_n_rt_rp_r_c_modf l_e_st_eq_landau_n_1 f)). Time Defined. (* constant 6295 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t5 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8288_prop1 l_e_st_eq_landau_n_1 f). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_modf l_e_st_eq_landau_n_1 f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t3 f) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t4 f)). Time Defined. (* constant 6296 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t6 : l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 l_e_st_eq_landau_n_1. exact (fun (u:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8288_t5 u). Time Defined. (* constant 6297 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t7 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))). Time Defined. (* constant 6298 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_lf : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8288_t7 x p f) f))). Time Defined. (* constant 6299 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t8 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) x f))). Time Defined. (* constant 6300 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_m : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_real))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_mod (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))). Time Defined. (* constant 6301 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t9 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_ismod (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t8 x p f)))). Time Defined. (* constant 6302 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t10 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz268 (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))). Time Defined. (* constant 6303 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t11 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t9 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t10 x p f)))). Time Defined. (* constant 6304 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx1 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_t11 x p f)))). Time Defined. (* constant 6305 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_lmf : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8288_t7 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)))). Time Defined. (* constant 6306 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t13 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lmf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) x (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)))). Time Defined. (* constant 6307 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t14 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lmf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lmf x p f)) (p (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))))). Time Defined. (* constant 6308 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t15 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lmf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lmf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t14 x p f)))). Time Defined. (* constant 6309 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t16 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_tsis12a (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 6310 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t17 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_ismn2 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz176b l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))))). Time Defined. (* constant 6311 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t18 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f))) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f))) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts02 (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 6312 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t19 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f))) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_t17 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t18 x p f)))). Time Defined. (* constant 6313 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t20 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lmf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t13 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t15 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t16 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t19 x p f)))). Time Defined. (* constant 6314 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t21 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8288_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_modf (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8288_lf x p f))) (l_e_st_eq_landau_n_rt_rp_r_c_8288_m x p f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t12 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t20 x p f)))). Time Defined. (* constant 6315 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t21a : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8288_t21 x p u))). Time Defined. (* constant 6316 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_t22 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x), l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 (l_e_st_eq_landau_n_suc x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 x) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 t) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_rt_rp_r_c_8288_t21a x p) (l_e_st_eq_landau_n_satz4a x))). Time Defined. (* constant 6317 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz288 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_prod x f)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_modf x f)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 t) l_e_st_eq_landau_n_rt_rp_r_c_8288_t6 (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 t) => l_e_st_eq_landau_n_rt_rp_r_c_8288_t22 t u)) x f)). Time Defined. (* constant 6318 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x f) l_e_st_eq_landau_n_rt_rp_r_c_0c)). Time Defined. (* constant 6319 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_some (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c))). Time Defined. (* constant 6320 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_prop3 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), Prop)). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_iff (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 x f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 x f))). Time Defined. (* constant 6321 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 : (forall (x:l_e_st_eq_landau_n_nat), Prop). exact (fun (x:l_e_st_eq_landau_n_nat) => (forall (u:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop3 x u)). Time Defined. (* constant 6322 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t1 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) f). Time Defined. (* constant 6323 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t2 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f), l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t1 f) p)). Time Defined. (* constant 6324 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t3 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f)). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f) => l_somei (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t2 f p))). Time Defined. (* constant 6325 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t4 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f u) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_is (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) u (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1))))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f u) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_singlet_th1 u)))). Time Defined. (* constant 6326 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t5 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f), (forall (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f u) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f)))). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f) => (fun (u:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f u) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 f) (f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1)) (f u) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_8289_t1 f) (l_e_isf (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) l_e_st_eq_landau_n_rt_rp_r_c_cx f (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1) u (l_e_symis (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) u (l_e_st_eq_landau_n_xout l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t4 f p u i))) i)))). Time Defined. (* constant 6327 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t6 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f)). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f) => l_someapp (l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) p (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f) (fun (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t5 f p t u)))). Time Defined. (* constant 6328 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t7 : (forall (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop3 l_e_st_eq_landau_n_1 f). exact (fun (f:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_iffi (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 l_e_st_eq_landau_n_1 f) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t3 f t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 l_e_st_eq_landau_n_1 f) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t6 f t)). Time Defined. (* constant 6329 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t8 : l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 l_e_st_eq_landau_n_1. exact (fun (u:(forall (t:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t7 u). Time Defined. (* constant 6330 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t9 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_lessis x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_lessisi1 x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_satz18a x l_e_st_eq_landau_n_1)))). Time Defined. (* constant 6331 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_lf : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8289_t9 x p f) f))). Time Defined. (* constant 6332 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t10 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) x f))). Time Defined. (* constant 6333 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t11 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t10 x p f) q)))). Time Defined. (* constant 6334 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t12 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), l_or (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => l_e_st_eq_landau_n_rt_rp_r_c_satz221c (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t11 x p f q))))). Time Defined. (* constant 6335 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t13 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_iff_th3 (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (p (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) i))))). Time Defined. (* constant 6336 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t14 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_1to x), (forall (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_somei (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8289_t9 x p f) n) j))))))). Time Defined. (* constant 6337 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t15 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_someapp (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t13 x p f q i) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (fun (t:l_e_st_eq_landau_n_1to x) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t14 x p f q i t u))))))). Time Defined. (* constant 6338 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t16 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_somei (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) i))))). Time Defined. (* constant 6339 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t17 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => l_orapp (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t12 x p f q) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t15 x p f q t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t16 x p f q t))))). Time Defined. (* constant 6340 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t18 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))), l_e_st_eq_landau_n_rt_rp_r_c_is (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) l_e_st_eq_landau_n_rt_rp_r_c_0c (f n) (l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_c_cx f n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) j) i))))))). Time Defined. (* constant 6341 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t20 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (j:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) => l_e_st_eq_landau_n_rt_rp_r_c_satz221b (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t18 x p f q n i j)))))))). Time Defined. (* constant 6342 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_nat))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_st_eq_landau_n_rt_rp_r_c_inn (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n))))))). Time Defined. (* constant 6343 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t21 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_lessisi3 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) j)))))))). Time Defined. (* constant 6344 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t22 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), (forall (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => (fun (j:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n)) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t21 x p f q n i m j))))))))). Time Defined. (* constant 6345 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t23 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_nis (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_imp_th3 (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) m (fun (t:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t22 x p f q n i m t)))))))). Time Defined. (* constant 6346 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t24 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_ore1 (l_e_st_eq_landau_n_less (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t23 x p f q n i m)))))))). Time Defined. (* constant 6347 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t25 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) x))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_st_eq_landau_n_satz26 x (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t24 x p f q n i m)))))))). Time Defined. (* constant 6348 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_1to x))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_st_eq_landau_n_outn x (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t25 x p f q n i m)))))))). Time Defined. (* constant 6349 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t26 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_st_eq_landau_n_isinoutn x (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t25 x p f q n i m)))))))). Time Defined. (* constant 6350 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t27 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8289_t9 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_st_eq_landau_n_isoutni (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n) (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m)) (l_e_st_eq_landau_n_trlessis (l_e_st_eq_landau_n_rt_rp_r_c_inn x (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m)) x (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_1top x (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m)) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t9 x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t26 x p f q n i m)))))))). Time Defined. (* constant 6351 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t28 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8289_t9 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_tris (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_outn (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n1 x p f q n i m) (l_e_st_eq_landau_n_1top (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n)) (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8289_t9 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m)) (l_e_st_eq_landau_n_isoutinn (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) n) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t27 x p f q n i m)))))))). Time Defined. (* constant 6352 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t29 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_rt_rp_r_c_is (f n) (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m))))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_isf (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) l_e_st_eq_landau_n_rt_rp_r_c_cx f n (l_e_st_eq_landau_n_left1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_8289_t9 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m)) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t28 x p f q n i m)))))))). Time Defined. (* constant 6353 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t30 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m)) l_e_st_eq_landau_n_rt_rp_r_c_0c (f n) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t29 x p f q n i m) i))))))). Time Defined. (* constant 6354 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t31 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_somei (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_8289_n2 x p f q n i m) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t30 x p f q n i m)))))))). Time Defined. (* constant 6355 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t32 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_iff_th4 (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (p (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t31 x p f q n i m)))))))). Time Defined. (* constant 6356 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t34 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (m:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_st_eq_landau_n_rt_rp_r_c_satz221a (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t32 x p f q n i m)))))))). Time Defined. (* constant 6357 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t35 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_imp_th1 (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) l_e_st_eq_landau_n_rt_rp_r_c_0c) (fun (t:l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t20 x p f q n i t) (fun (t:l_not (l_e_is (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) n (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t34 x p f q n i t))))))). Time Defined. (* constant 6358 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t36 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), (forall (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)))))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => (fun (n:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod x (l_e_st_eq_landau_n_rt_rp_r_c_8289_lf x p f)) (f (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)))) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_8289_t10 x p f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t35 x p f q n i))))))). Time Defined. (* constant 6359 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t37 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => l_someapp (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) q (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t36 x p f q t u)))))). Time Defined. (* constant 6360 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t38 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), (forall (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop3 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_iffi (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t17 x p f t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) f) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t37 x p f t)))). Time Defined. (* constant 6361 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t39 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t38 x p u))). Time Defined. (* constant 6362 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t40 : (forall (x:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 (l_e_st_eq_landau_n_suc x))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 x) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 t) (l_e_st_eq_landau_n_pl x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_suc x) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t39 x p) (l_e_st_eq_landau_n_satz4a x))). Time Defined. (* constant 6363 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz289 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), l_iff (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x f) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_some (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 t) l_e_st_eq_landau_n_rt_rp_r_c_8289_t8 (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 t) => l_e_st_eq_landau_n_rt_rp_r_c_8289_t40 t u)) x f)). Time Defined. (* constant 6364 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz289a : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x f) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_some (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x f) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_iff_th3 (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 x f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 x f) (l_e_st_eq_landau_n_rt_rp_r_c_satz289 x f) i))). Time Defined. (* constant 6365 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_t41 : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 x f)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_somei (l_e_st_eq_landau_n_1to x) (fun (t:l_e_st_eq_landau_n_1to x) => l_e_st_eq_landau_n_rt_rp_r_c_is (f t) l_e_st_eq_landau_n_rt_rp_r_c_0c) n i)))). Time Defined. (* constant 6366 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz289b : (forall (x:l_e_st_eq_landau_n_nat), (forall (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)), (forall (n:l_e_st_eq_landau_n_1to x), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod x f) l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_nat) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to x), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (n:l_e_st_eq_landau_n_1to x) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (f n) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_iff_th4 (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop1 x f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_prop2 x f) (l_e_st_eq_landau_n_rt_rp_r_c_satz289 x f) (l_e_st_eq_landau_n_rt_rp_r_c_8289_t41 x f n i))))). Time Defined. (* constant 6367 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_natrl m))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_posintnatrl m p mi))))). Time Defined. (* constant 6368 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_nat))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_ntofrl m (l_e_st_eq_landau_n_rt_rp_r_c_v9_t1 x m mi o p)))))). Time Defined. (* constant 6369 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_cx))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => x)))))). Time Defined. (* constant 6370 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (np:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_nat)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (np:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi1 ox mp)))))))))))). Time Defined. (* constant 6371 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (np:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_nat)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (np:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 y n ni1 oy np)))))))))))). Time Defined. (* constant 6372 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (np:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (np:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_isrlent m (l_e_st_eq_landau_n_rt_rp_r_c_v9_t1 x m mi1 ox mp) n (l_e_st_eq_landau_n_rt_rp_r_c_v9_t1 y n ni1 oy np) j)))))))))))). Time Defined. (* constant 6373 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (np:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (np:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t2 x y m n i j mi1 ni1 ox oy mp np))))))))))))). Time Defined. (* constant 6374 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (np:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t3 x y m n i j mi1 ni1 ox oy mp np) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np)) => y))) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 y n ni1 oy np))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (np:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np)) => y) (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t2 x y m n i j mi1 ni1 ox oy mp np))))))))))))). Time Defined. (* constant 6375 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (np:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_is (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np)), l_e_st_eq_landau_n_rt_rp_r_c_cx) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np)) => x) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t3 x y m n i j mi1 ni1 ox oy mp np) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np)) => y)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (np:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_fisi (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np)) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np)) => x) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t3 x y m n i j mi1 ni1 ox oy mp np) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np)) => y)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np)) => i))))))))))))). Time Defined. (* constant 6376 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (np:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi1 ox mp) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t3 x y m n i j mi1 ni1 ox oy mp np) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np)) => y))))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (np:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_isf (forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np)), l_e_st_eq_landau_n_rt_rp_r_c_cx) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (u:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) u) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np)) => x) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t3 x y m n i j mi1 ni1 ox oy mp np) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np)) => y)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t5 x y m n i j mi1 ni1 ox oy mp np))))))))))))). Time Defined. (* constant 6377 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (np:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi1 ox mp) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 y n ni1 oy np))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (np:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi1 ox mp) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_m0 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t3 x y m n i j mi1 ni1 ox oy mp np) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_n0 x y m n i j mi1 ni1 ox oy mp np)) => y))) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 y n ni1 oy np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t6 x y m n i j mi1 ni1 ox oy mp np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t4 x y m n i j mi1 ni1 ox oy mp np))))))))))))). Time Defined. (* constant 6378 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (p1:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o p) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o p1))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (p1:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t7 x x m m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real m) mi mi o o p p1)))))). Time Defined. (* constant 6379 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_not (l_some (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => l_e_st_eq_landau_n_rt_rp_r_c_is ((fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => x) t) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_some_th5 (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => l_e_st_eq_landau_n_rt_rp_r_c_is ((fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => x) t) l_e_st_eq_landau_n_rt_rp_r_c_0c) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => n))))))). Time Defined. (* constant 6380 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o p) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o p) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_some (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => l_e_st_eq_landau_n_rt_rp_r_c_is ((fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => x) t) l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t9 x m mi o p n) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o p) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz289a (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p) (fun (u:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_v9_m1 x m mi o p)) => x) t))))))). Time Defined. (* constant 6381 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_abs m)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_st_eq_landau_n_rt_rp_r_intabs m mi))). Time Defined. (* constant 6382 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_satz166b m n))))). Time Defined. (* constant 6383 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n)))))). Time Defined. (* constant 6384 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m) o (l_e_st_eq_landau_n_rt_rp_r_nnotp m n)))))). Time Defined. (* constant 6385 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t15 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t10 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t14 x m mi o n)))))). Time Defined. (* constant 6386 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_cx))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t15 x m mi o n)))))). Time Defined. (* constant 6387 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_pwm : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi1) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi1 ox nm) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi1 ox nm))))))))))))). Time Defined. (* constant 6388 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_pwn : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 y (l_e_st_eq_landau_n_rt_rp_r_abs n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 y n ni1) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 y n ni1 oy nn) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 y n ni1 oy nn))))))))))))). Time Defined. (* constant 6389 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t16 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pwm x y m n i j mi1 ni1 ox oy nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pwn x y m n i j mi1 ni1 ox oy nm nn))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t7 x y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_abs n) i (l_e_st_eq_landau_n_rt_rp_r_isabs m n j) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi1) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 y n ni1) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi1 ox nm) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 y n ni1 oy nn) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi1 ox nm) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 y n ni1 oy nn))))))))))))). Time Defined. (* constant 6390 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t17 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi1 ox nm) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 y n ni1 oy nn))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_isov2 (l_e_st_eq_landau_n_rt_rp_r_c_v9_pwm x y m n i j mi1 ni1 ox oy nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pwn x y m n i j mi1 ni1 ox oy nm nn) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_v9_t16 x y m n i j mi1 ni1 ox oy nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t15 x m mi1 ox nm) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t15 y n ni1 oy nn))))))))))))). Time Defined. (* constant 6391 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t18 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (n1:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o n1))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (n1:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t17 x x m m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real m) mi mi o o n n1)))))). Time Defined. (* constant 6392 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_r_ite (l_e_st_eq_landau_n_rt_rp_r_neg m) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_st_eq_landau_n_rt_rp_r_c_1c) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t18 x m mi o t u)) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_1c)))))). Time Defined. (* constant 6393 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t19 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_r_itet (l_e_st_eq_landau_n_rt_rp_r_neg m) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_st_eq_landau_n_rt_rp_r_c_1c) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t18 x m mi o t u)) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_1c)) n))))). Time Defined. (* constant 6394 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t20 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_1c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_r_itef (l_e_st_eq_landau_n_rt_rp_r_neg m) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_st_eq_landau_n_rt_rp_r_c_1c) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t18 x m mi o t u)) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_1c)) nn))))). Time Defined. (* constant 6395 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t21 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_neg n))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_neg t) m n nm j))))))))))). Time Defined. (* constant 6396 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t22 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi1 ox nm)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t19 x m mi1 ox nm))))))))))). Time Defined. (* constant 6397 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t23 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t21 x y m n i j mi1 ni1 ox oy nm)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t21 x y m n i j mi1 ni1 ox oy nm)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t19 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t21 x y m n i j mi1 ni1 ox oy nm))))))))))))). Time Defined. (* constant 6398 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t24 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi1 ox nm) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t21 x y m n i j mi1 ni1 ox oy nm)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t22 x y m n i j mi1 ni1 ox oy nm) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t17 x y m n i j mi1 ni1 ox oy nm (l_e_st_eq_landau_n_rt_rp_r_c_v9_t21 x y m n i j mi1 ni1 ox oy nm)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t23 x y m n i j mi1 ni1 ox oy nm)))))))))))). Time Defined. (* constant 6399 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t25 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)), l_not (l_e_st_eq_landau_n_rt_rp_r_neg n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_not (l_e_st_eq_landau_n_rt_rp_r_neg t)) m n nn j))))))))))). Time Defined. (* constant 6400 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t26 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) l_e_st_eq_landau_n_rt_rp_r_c_1c))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t20 x m mi1 ox nn))))))))))). Time Defined. (* constant 6401 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t27 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy) l_e_st_eq_landau_n_rt_rp_r_c_1c))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t20 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t25 x y m n i j mi1 ni1 ox oy nn)))))))))))). Time Defined. (* constant 6402 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t28 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (nn:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_v9_t26 x y m n i j mi1 ni1 ox oy nn) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t27 x y m n i j mi1 ni1 ox oy nn)))))))))))). Time Defined. (* constant 6403 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t29 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_neg m) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t24 x y m n i j mi1 ni1 ox oy t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_neg m)) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t28 x y m n i j mi1 ni1 ox oy t))))))))))). Time Defined. (* constant 6404 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_pw : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_r_ite (l_e_st_eq_landau_n_rt_rp_r_pos m) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t8 x m mi o t u)) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o))))))). Time Defined. (* constant 6405 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t30 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o p)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_r_itet (l_e_st_eq_landau_n_rt_rp_r_pos m) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t8 x m mi o t u)) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o))) p))))). Time Defined. (* constant 6406 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t31 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_r_itef (l_e_st_eq_landau_n_rt_rp_r_pos m) l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t8 x m mi o t u)) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (u:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o))) n))))). Time Defined. (* constant 6407 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t32 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_1c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_v9_t31 x m mi o (l_e_st_eq_landau_n_rt_rp_r_0notp m i)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t20 x m mi o (l_e_st_eq_landau_n_rt_rp_r_0notn m i))))))). Time Defined. (* constant 6408 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t33 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t31 x m mi o (l_e_st_eq_landau_n_rt_rp_r_nnotp m n)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t19 x m mi o n)))))). Time Defined. (* constant 6409 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_posexp : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_ntofrl m (l_e_st_eq_landau_n_rt_rp_r_posintnatrl m p mi)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_ntofrl m (l_e_st_eq_landau_n_rt_rp_r_posintnatrl m p mi))) => x))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t30 x m mi o p))))). Time Defined. (* constant 6410 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_notis_th2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi o p) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t10 x m mi o p n) (l_e_st_eq_landau_n_rt_rp_r_c_posexp x m mi o p))))))). Time Defined. (* constant 6411 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_0exp : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_1c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t32 x m mi o i))))). Time Defined. (* constant 6412 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t14 x m mi o n))))). Time Defined. (* constant 6413 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n))))). Time Defined. (* constant 6414 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t34 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t30 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n)))))). Time Defined. (* constant 6415 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t35 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t14 x m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_isov2 (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_v9_t34 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t11 x m mi) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t13 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t12 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t14 x m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t15 x m mi o n)))))). Time Defined. (* constant 6416 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_negexp : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_intabs m mi) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_intabs m mi) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_satz166b m n) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x m mi o n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_intabs m mi) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_intabs m mi) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_satz166b m n) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw2 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t33 x m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t35 x m mi o n)))))). Time Defined. (* constant 6417 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t36 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_pos n))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pos t) m n mp j))))))))))). Time Defined. (* constant 6418 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t37 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi1 ox mp)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t30 x m mi1 ox mp))))))))))). Time Defined. (* constant 6419 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t38 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t36 x y m n i j mi1 ni1 ox oy mp)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t36 x y m n i j mi1 ni1 ox oy mp)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t30 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t36 x y m n i j mi1 ni1 ox oy mp))))))))))))). Time Defined. (* constant 6420 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t39 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (mp:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (mp:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 x m mi1 ox mp) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw1 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t36 x y m n i j mi1 ni1 ox oy mp)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t37 x y m n i j mi1 ni1 ox oy mp) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t7 x y m n i j mi1 ni1 ox oy mp (l_e_st_eq_landau_n_rt_rp_r_c_v9_t36 x y m n i j mi1 ni1 ox oy mp)) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t38 x y m n i j mi1 ni1 ox oy mp)))))))))))). Time Defined. (* constant 6421 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t40 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (np:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_not (l_e_st_eq_landau_n_rt_rp_r_pos n)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (np:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_not (l_e_st_eq_landau_n_rt_rp_r_pos t)) m n np j))))))))))). Time Defined. (* constant 6422 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t41 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (np:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (np:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t31 x m mi1 ox np))))))))))). Time Defined. (* constant 6423 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t42 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (np:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (np:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t31 y n ni1 oy (l_e_st_eq_landau_n_rt_rp_r_c_v9_t40 x y m n i j mi1 ni1 ox oy np))))))))))))). Time Defined. (* constant 6424 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v9_t43 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), (forall (np:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => (fun (np:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_v9_pw3 y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t41 x y m n i j mi1 ni1 ox oy np) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t29 x y m n i j mi1 ni1 ox oy) (l_e_st_eq_landau_n_rt_rp_r_c_v9_t42 x y m n i j mi1 ni1 ox oy np)))))))))))). Time Defined. (* constant 6425 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ispw12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (j:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (j:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi1:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni1:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi1 ox) (l_e_st_eq_landau_n_rt_rp_r_c_pw y n ni1 oy)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t39 x y m n i j mi1 ni1 ox oy t) (fun (t:l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_st_eq_landau_n_rt_rp_r_c_v9_t43 x y m n i j mi1 ni1 ox oy t))))))))))). Time Defined. (* constant 6426 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ispw1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi ox) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi oy)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ox:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (oy:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw12 x y m m i (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real m) mi mi ox oy))))))). Time Defined. (* constant 6427 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ispw2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m n), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (om:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (on:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi om) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni on))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m n) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (om:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (on:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw12 x x m n (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) i mi ni om on)))))))). Time Defined. (* constant 6428 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9290_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x m mi o p n)))))). Time Defined. (* constant 6429 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9290_t2 : (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0). exact (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_re l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_re l_e_st_eq_landau_n_rt_rp_r_c_0c) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_isre l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_iscere l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c i) (l_e_st_eq_landau_n_rt_rp_r_c_reis l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)). Time Defined. (* constant 6430 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_1not0 : l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c. exact (l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_c_is l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_is l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pnot0 l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_pos1) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_9290_t2 t)). Time Defined. (* constant 6431 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9290_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_notis_th2 l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_1not0 (l_e_st_eq_landau_n_rt_rp_r_c_0exp x m mi o i))))))). Time Defined. (* constant 6432 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_intabs m mi) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x m mi o nm))))))). Time Defined. (* constant 6433 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9290_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_intabs m mi) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x m mi o nm) (l_e_st_eq_landau_n_rt_rp_r_satz166b m nm) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x m mi o nm))))))). Time Defined. (* constant 6434 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9290_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm)) l_e_st_eq_landau_n_rt_rp_r_c_1c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm) (l_e_st_eq_landau_n_rt_rp_r_c_9290_t5 x m mi o n nm)) (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm) (l_e_st_eq_landau_n_rt_rp_r_c_9290_t5 x m mi o n nm)) (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm) (l_e_st_eq_landau_n_rt_rp_r_c_negexp x m mi o nm)) (l_e_st_eq_landau_n_rt_rp_r_c_satz229e l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm) (l_e_st_eq_landau_n_rt_rp_r_c_9290_t5 x m mi o n nm)))))))). Time Defined. (* constant 6435 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9290_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_notis_th2 l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm)) l_e_st_eq_landau_n_rt_rp_r_c_1not0 (l_e_st_eq_landau_n_rt_rp_r_c_9290_t6 x m mi o n nm))))))). Time Defined. (* constant 6436 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9290_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm)) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_9290_t7 x m mi o n nm) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz221a (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_9290_p0 x m mi o n nm) t))))))). Time Defined. (* constant 6437 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz290 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_rapp m (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_9290_t1 x m mi o n t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_9290_t4 x m mi o n t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_9290_t8 x m mi o n t)))))). Time Defined. (* constant 6438 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma291 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_1rl)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_1rl) l_e_st_eq_landau_n_rt_rp_r_pos1). Time Defined. (* constant 6439 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9291_1a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_nat). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_ntofrl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_posintnatrl l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_pos1 l_e_st_eq_landau_n_rt_rp_r_intrl1)). Time Defined. (* constant 6440 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9291_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 x)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_posexp x l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 x) l_e_st_eq_landau_n_rt_rp_r_pos1). Time Defined. (* constant 6441 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9291_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_is l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_ntofrl l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_natrl1) (l_e_st_eq_landau_n_rt_rp_r_ntofrl l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_posintnatrl l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_pos1 l_e_st_eq_landau_n_rt_rp_r_intrl1)) (l_e_st_eq_landau_n_rt_rp_r_isntrl1 l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_isrlent l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_natrl1 l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_posintnatrl l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_pos1 l_e_st_eq_landau_n_rt_rp_r_intrl1) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_1rl))). Time Defined. (* constant 6442 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9291_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_lessis l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_lessisi2 l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) (l_e_st_eq_landau_n_rt_rp_r_c_9291_t2 x)). Time Defined. (* constant 6443 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9291_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_9291_t3 x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x))) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_9291_t2 x)). Time Defined. (* constant 6444 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9291_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_9291_t3 x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x))) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_st_eq_landau_n_rt_rp_r_c_satz277 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_9291_t3 x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x))). Time Defined. (* constant 6445 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9291_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x)) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris1 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x)) x (l_e_st_eq_landau_n_rt_rp_r_c_prod l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) l_e_st_eq_landau_n_1 (l_e_st_eq_landau_n_rt_rp_r_c_9291_t3 x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x))) (l_e_st_eq_landau_n_rt_rp_r_c_9291_t4 x) (l_e_st_eq_landau_n_rt_rp_r_c_9291_t5 x)). Time Defined. (* constant 6446 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz291 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 x)) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 x)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9291_1a x)) => x)) x (l_e_st_eq_landau_n_rt_rp_r_c_9291_t1 x) (l_e_st_eq_landau_n_rt_rp_r_c_9291_t6 x)). Time Defined. (* constant 6447 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)), l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) a)))))). Time Defined. (* constant 6448 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)), l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) a)))))). Time Defined. (* constant 6449 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) => l_e_st_eq_landau_n_rt_rp_r_c_satz221d x y (l_e_st_eq_landau_n_rt_rp_r_c_9292_t1 x y m mi o a) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t2 x y m mi o a))))))). Time Defined. (* constant 6450 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma292a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_or_th7 (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) o (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) => l_e_st_eq_landau_n_rt_rp_r_c_9292_t1 x y m mi o t)))))). Time Defined. (* constant 6451 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma292b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_or_th7 (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) o (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) => l_e_st_eq_landau_n_rt_rp_r_c_9292_t2 x y m mi o t)))))). Time Defined. (* constant 6452 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma292c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_or_th7 (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) o (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) => l_e_st_eq_landau_n_rt_rp_r_c_9292_t3 x y m mi o t)))))). Time Defined. (* constant 6453 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_nr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt n)). Time Defined. (* constant 6454 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_natintrl (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natrli n))). Time Defined. (* constant 6455 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n)) (l_e_st_eq_landau_n_rt_rp_r_natpos (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natrli n)))). Time Defined. (* constant 6456 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_cx)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t4 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t5 x n))). Time Defined. (* constant 6457 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x l_e_st_eq_landau_n_1) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 x)) x (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9292_t5 x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 x)) (l_e_st_eq_landau_n_rt_rp_r_c_satz291 x)). Time Defined. (* constant 6458 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_posintnatrl (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natpos (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natrli n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t4 x n)))). Time Defined. (* constant 6459 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is n (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat n (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natrli n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n) (l_e_st_eq_landau_n_rt_rp_r_isntrl1 n) (l_e_st_eq_landau_n_rt_rp_r_isrlent (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natrli n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_posintnatrl (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natpos (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natrli n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t4 x n)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n))))). Time Defined. (* constant 6460 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lessis n (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessisi2 n (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t7 x n))). Time Defined. (* constant 6461 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n)) => x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_posexp x (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t4 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t5 x n) (l_e_st_eq_landau_n_rt_rp_r_natpos (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x n) (l_e_st_eq_landau_n_rt_rp_r_natrli n)))). Time Defined. (* constant 6462 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod n (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n) n (l_e_st_eq_landau_n_rt_rp_r_c_9292_t8 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n)) => x))) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n)) => x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n)) => x) n (l_e_st_eq_landau_n_rt_rp_r_c_9292_t7 x n))). Time Defined. (* constant 6463 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_prod n (fun (t:l_e_st_eq_landau_n_1to n) => x)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_prod n (fun (t:l_e_st_eq_landau_n_1to n) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n0 x n)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t9 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t10 x n))). Time Defined. (* constant 6464 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_pl n l_e_st_eq_landau_n_1)). Time Defined. (* constant 6465 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_lessis n (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_lessisi1 n (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n) (l_e_st_eq_landau_n_satz18a n l_e_st_eq_landau_n_1))). Time Defined. (* constant 6466 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod n (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n) n (l_e_st_eq_landau_n_rt_rp_r_c_9292_t12 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) => x))) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_satz278 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) n (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) => x))). Time Defined. (* constant 6467 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod n (fun (t:l_e_st_eq_landau_n_1to n) => x)) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_prod n (fun (t:l_e_st_eq_landau_n_1to n) => x)) x (l_e_st_eq_landau_n_rt_rp_r_c_9292_t11 x n))). Time Defined. (* constant 6468 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t15 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod n (fun (t:l_e_st_eq_landau_n_1to n) => x)) x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t13 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t14 x n))). Time Defined. (* constant 6469 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t16 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t11 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t15 x n))). Time Defined. (* constant 6470 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), Prop))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n))))). Time Defined. (* constant 6471 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t17 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x l_e_st_eq_landau_n_1) x (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y l_e_st_eq_landau_n_1) y (l_e_st_eq_landau_n_rt_rp_r_c_9292_t6 x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t6 y))). Time Defined. (* constant 6472 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t18 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y l_e_st_eq_landau_n_1)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t6 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t17 x y))). Time Defined. (* constant 6473 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t19 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n) => l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) p)))). Time Defined. (* constant 6474 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t20 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts x (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_comts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) x)) (l_e_st_eq_landau_n_rt_rp_r_c_assts2 (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)))))). Time Defined. (* constant 6475 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t21 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) x) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) y) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) y)) (l_e_st_eq_landau_n_rt_rp_r_c_assts2 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) x y) (l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) y (l_e_st_eq_landau_n_rt_rp_r_c_9292_t20 x y n p)) (l_e_st_eq_landau_n_rt_rp_r_c_assts1 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) y))))). Time Defined. (* constant 6476 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t22 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts x y)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) y)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t16 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t19 x y n p) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t21 x y n p))))). Time Defined. (* constant 6477 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t23 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) y)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t16 x n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t16 y n))))). Time Defined. (* constant 6478 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t24 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n), l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x n) x) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y n) y)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t22 x y n p) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t23 x y n p))))). Time Defined. (* constant 6479 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t25 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n), l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y (l_e_st_eq_landau_n_suc n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y t) (l_e_st_eq_landau_n_rt_rp_r_c_9292_n1 x n) (l_e_st_eq_landau_n_suc n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t24 x y n p) (l_e_st_eq_landau_n_satz4a n))))). Time Defined. (* constant 6480 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t26 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y n))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y t) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t18 x y) (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_9292_prop1 x y t) => l_e_st_eq_landau_n_rt_rp_r_c_9292_t25 x y t u)) n))). Time Defined. (* constant 6481 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_prop2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), Prop))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o)))))))). Time Defined. (* constant 6482 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t28 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_natrl m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_posintnatrl m p mi)))))). Time Defined. (* constant 6483 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_nat)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_ntofrl m (l_e_st_eq_landau_n_rt_rp_r_c_9292_t28 x y m mi o p))))))). Time Defined. (* constant 6484 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t29 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_isrlnt1 m (l_e_st_eq_landau_n_rt_rp_r_c_9292_t28 x y m mi o p))))))). Time Defined. (* constant 6485 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t30 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_isrlnt2 m (l_e_st_eq_landau_n_rt_rp_r_c_9292_t28 x y m mi o p))))))). Time Defined. (* constant 6486 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t31 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t29 x y m mi o p) mi (l_e_st_eq_landau_n_rt_rp_r_c_9292_t4 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t5 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)))))))). Time Defined. (* constant 6487 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t32 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) m (l_e_st_eq_landau_n_rt_rp_r_c_9292_t30 x y m mi o p) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t4 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) mi (l_e_st_eq_landau_n_rt_rp_r_c_9292_t5 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw2 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_nr x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) m (l_e_st_eq_landau_n_rt_rp_r_c_9292_t30 x y m mi o p) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t4 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) mi (l_e_st_eq_landau_n_rt_rp_r_c_9292_t5 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o)))))))). Time Defined. (* constant 6488 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t33 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_9292_prop2 x y m mi o)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 x (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_p0 y (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t31 x y m mi o p) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t26 x y (l_e_st_eq_landau_n_rt_rp_r_c_9292_m0 x y m mi o p)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t32 x y m mi o p))))))). Time Defined. (* constant 6489 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t34 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))) l_e_st_eq_landau_n_rt_rp_r_c_1c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_0exp x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o) i) (l_e_st_eq_landau_n_rt_rp_r_c_0exp y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o) i)) (l_e_st_eq_landau_n_rt_rp_r_c_satz222 l_e_st_eq_landau_n_rt_rp_r_c_1c))))))). Time Defined. (* constant 6490 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t35 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_9292_prop2 x y m mi o)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_0exp (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o) i) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t34 x y m mi o i))))))). Time Defined. (* constant 6491 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_abs m))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_intabs m mi)))))). Time Defined. (* constant 6492 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t37 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_ori2 (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)) (l_e_st_eq_landau_n_rt_rp_r_satz166b m n))))))). Time Defined. (* constant 6493 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t38 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t37 x y m mi o n))))))). Time Defined. (* constant 6494 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t39 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t37 x y m mi o n))))))). Time Defined. (* constant 6495 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t40 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t37 x y m mi o n))))))). Time Defined. (* constant 6496 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o) n)))))). Time Defined. (* constant 6497 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o) n)))))). Time Defined. (* constant 6498 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_abs m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o) n)))))). Time Defined. (* constant 6499 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t44 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t40 x y m mi o n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs m)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t40 x y m mi o n))))))). Time Defined. (* constant 6500 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t45 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t40 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t38 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t39 x y m mi o n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_9292_t33 x y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t37 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_satz166b m n))))))). Time Defined. (* constant 6501 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t46 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t38 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t39 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t38 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t39 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs m)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t38 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw2 y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs m)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t39 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)))))))). Time Defined. (* constant 6502 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t47 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t40 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t38 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t39 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t44 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t45 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t46 x y m mi o n))))))). Time Defined. (* constant 6503 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t48 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_satz166b m n) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o) n))))))). Time Defined. (* constant 6504 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t49 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_satz166b m n) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o) n))))))). Time Defined. (* constant 6505 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t50 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_satz166b m n) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o) n))))))). Time Defined. (* constant 6506 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_satz221d (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t48 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t49 x y m mi o n))))))). Time Defined. (* constant 6507 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t52 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t50 x y m mi o n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_negexp (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o) n)))))). Time Defined. (* constant 6508 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t53 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t50 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 x y m mi o n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_isov12 l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_satz222a l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t47 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t50 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 x y m mi o n))))))). Time Defined. (* constant 6509 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t54 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 x y m mi o n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t43 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t50 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t52 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t53 x y m mi o n))))))). Time Defined. (* constant 6510 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t55 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t48 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t49 x y m mi o n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t48 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t49 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_negexp x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o) n) (l_e_st_eq_landau_n_rt_rp_r_c_negexp y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o) n))))))). Time Defined. (* constant 6511 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t56 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t48 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t49 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 x y m mi o n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_satz247 l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t48 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t49 x y m mi o n))))))). Time Defined. (* constant 6512 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t57 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 x y m mi o n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t48 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t49 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t55 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t56 x y m mi o n))))))). Time Defined. (* constant 6513 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9292_t58 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_9292_prop2 x y m mi o)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t41 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y (l_e_st_eq_landau_n_rt_rp_r_abs m) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t36 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t42 x y m mi o n))) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t51 x y m mi o n)) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t54 x y m mi o n) (l_e_st_eq_landau_n_rt_rp_r_c_9292_t57 x y m mi o n))))))). Time Defined. (* constant 6514 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz292 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts x y) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a x y m mi o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw y m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b x y m mi o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis y l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => l_e_st_eq_landau_n_rt_rp_r_rapp m (l_e_st_eq_landau_n_rt_rp_r_c_9292_prop2 x y m mi o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_9292_t33 x y m mi o t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_9292_t35 x y m mi o t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_9292_t58 x y m mi o t)))))). Time Defined. (* constant 6515 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma293 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m) l_e_st_eq_landau_n_rt_rp_r_c_1not0). Time Defined. (* constant 6516 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_or (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m))). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_ori1 (l_and (l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c)) (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_andi (l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_nis l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_0c) l_e_st_eq_landau_n_rt_rp_r_c_1not0 l_e_st_eq_landau_n_rt_rp_r_c_1not0))). Time Defined. (* constant 6517 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_1m : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_cx)). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma293 m))). Time Defined. (* constant 6518 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_t2 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi))). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_st_eq_landau_n_rt_rp_r_c_satz222 (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi))). Time Defined. (* constant 6519 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_t3 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))))). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_st_eq_landau_n_rt_rp_r_c_ispw1 l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) m (l_e_st_eq_landau_n_rt_rp_r_c_satz222a l_e_st_eq_landau_n_rt_rp_r_c_1c) mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma293 m) (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi)))). Time Defined. (* constant 6520 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_t4 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))) (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi)))))). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_st_eq_landau_n_rt_rp_r_c_satz292 l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))). Time Defined. (* constant 6521 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_t5 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))) (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi)))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi)))). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_1c) mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma293 m)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx l_e_st_eq_landau_n_rt_rp_r_c_1c) mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma293 m)))). Time Defined. (* constant 6522 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_t6 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi)))). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292c l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292a l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi))) (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma292b l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_9293_t1 m mi)))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi)) (l_e_st_eq_landau_n_rt_rp_r_c_9293_t2 m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_t3 m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_t4 m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_t5 m mi))). Time Defined. (* constant 6523 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_t7 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c)) l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c)) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_disttm2 (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_satz213b (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi)) (l_e_st_eq_landau_n_rt_rp_r_c_9293_t6 m mi))))). Time Defined. (* constant 6524 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9293_t8 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) l_e_st_eq_landau_n_rt_rp_r_c_0c)). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_ore2 (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_satz221c (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_9293_t7 m mi)) (l_e_st_eq_landau_n_rt_rp_r_c_satz290 l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma293 m) l_e_st_eq_landau_n_rt_rp_r_c_1not0))). Time Defined. (* constant 6525 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz293 : (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw l_e_st_eq_landau_n_rt_rp_r_c_1c m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma293 m)) l_e_st_eq_landau_n_rt_rp_r_c_1c)). exact (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => l_e_st_eq_landau_n_rt_rp_r_c_satz213a (l_e_st_eq_landau_n_rt_rp_r_c_9293_1m m mi) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_9293_t8 m mi))). Time Defined. (* constant 6526 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_pos m))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_ande1 (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n) a))))))). Time Defined. (* constant 6527 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_pos n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n) a))))))). Time Defined. (* constant 6528 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl m n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_pospl m n (l_e_st_eq_landau_n_rt_rp_r_c_9294_t1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t2 x m n mi ni o a)))))))). Time Defined. (* constant 6529 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma294a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_or_th8 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) (l_e_st_eq_landau_n_rt_rp_r_pos m) o (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t1 x m n mi ni o t))))))). Time Defined. (* constant 6530 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma294b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_or_th8 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) (l_e_st_eq_landau_n_rt_rp_r_pos n) o (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t2 x m n mi ni o t))))))). Time Defined. (* constant 6531 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma294c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl m n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_or_th8 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl m n)) o (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t3 x m n mi ni o t))))))). Time Defined. (* constant 6532 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), Prop)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)))))))). Time Defined. (* constant 6533 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_nat))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_ntofrl m (l_e_st_eq_landau_n_rt_rp_r_posintnatrl m (l_e_st_eq_landau_n_rt_rp_r_c_9294_t1 x m n mi ni o a) mi)))))))). Time Defined. (* constant 6534 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_nat))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_ntofrl n (l_e_st_eq_landau_n_rt_rp_r_posintnatrl n (l_e_st_eq_landau_n_rt_rp_r_c_9294_t2 x m n mi ni o a) ni)))))))). Time Defined. (* constant 6535 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) => x)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_posexp x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_c_posexp x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t2 x m n mi ni o a))))))))). Time Defined. (* constant 6536 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_nat))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_posintnatrl (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t3 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni))))))))). Time Defined. (* constant 6537 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a)) => x))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_posexp x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t3 x m n mi ni o a)))))))). Time Defined. (* constant 6538 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) (l_e_st_eq_landau_n_rt_rp_r_ispl12 m (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a)) n (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 m (l_e_st_eq_landau_n_rt_rp_r_posintnatrl m (l_e_st_eq_landau_n_rt_rp_r_c_9294_t1 x m n mi ni o a) mi)) (l_e_st_eq_landau_n_rt_rp_r_isrlnt1 n (l_e_st_eq_landau_n_rt_rp_r_posintnatrl n (l_e_st_eq_landau_n_rt_rp_r_c_9294_t2 x m n mi ni o a) ni))) (l_e_st_eq_landau_n_rt_rp_r_satzr155b (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))))))))). Time Defined. (* constant 6539 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_tris2 l_e_st_eq_landau_n_nat (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_ntofrl (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)))) (l_e_st_eq_landau_n_rt_rp_r_isntrl1 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) (l_e_st_eq_landau_n_rt_rp_r_isrlent (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_posintnatrl (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t3 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni)) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) (l_e_st_eq_landau_n_rt_rp_r_natrli (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t6 x m n mi ni o a))))))))). Time Defined. (* constant 6540 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_lessisi2 (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t7 x m n mi ni o a)))))))). Time Defined. (* constant 6541 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t8 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a)) => x))) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a)) => x))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_issmpr (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a)) => x) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t7 x m n mi ni o a)))))))). Time Defined. (* constant 6542 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) => x))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_p1 x m n mi ni o a)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t5 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t9 x m n mi ni o a)))))))). Time Defined. (* constant 6543 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_lessis (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_lessisi1 (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_satz18a (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))))))))). Time Defined. (* constant 6544 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_left l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t11 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) => x))) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a) (l_e_st_eq_landau_n_right l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) => x))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_satz281 (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_ts t u)) l_e_st_eq_landau_n_rt_rp_r_c_assocts (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) => x)))))))). Time Defined. (* constant 6545 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) => x)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a))) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) => x))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t10 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t12 x m n mi ni o a)))))))). Time Defined. (* constant 6546 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_m1 x m n mi ni o a)) => x)) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9294_n1 x m n mi ni o a)) => x))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t4 x m n mi ni o a) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t13 x m n mi ni o a)))))))). Time Defined. (* constant 6547 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t15 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_ore1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) o na))))))). Time Defined. (* constant 6548 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t16 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) (l_not (l_e_st_eq_landau_n_rt_rp_r_pos n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_or_th15 (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n) na))))))). Time Defined. (* constant 6549 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_am : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_real)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_abs m)))))). Time Defined. (* constant 6550 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_an : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_real)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_abs n)))))). Time Defined. (* constant 6551 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_ap : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_real)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_pl m n))))))). Time Defined. (* constant 6552 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_intabs m mi)))))). Time Defined. (* constant 6553 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_intabs n ni)))))). Time Defined. (* constant 6554 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_intabs (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni))))))). Time Defined. (* constant 6555 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t20 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_andi (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_satz166e m (l_e_st_eq_landau_n_rt_rp_r_nnot0 m nm)) (l_e_st_eq_landau_n_rt_rp_r_satz166e n (l_e_st_eq_landau_n_rt_rp_r_nnot0 n nn))))))))))). Time Defined. (* constant 6556 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t21 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t20 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6557 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) nm))))))))). Time Defined. (* constant 6558 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) nn))))))))). Time Defined. (* constant 6559 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t24 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t21 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6560 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t25 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t21 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6561 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t26 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t24 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t25 x m n mi ni o na nm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t24 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t25 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t24 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t25 x m n mi ni o na nm nn))))))))))). Time Defined. (* constant 6562 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t27 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t21 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6563 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t28 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t24 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t25 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t27 x m n mi ni o na nm nn))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t14 x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t21 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t20 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6564 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t29 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_m0 n)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl m n)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_ispl12 (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_m0 n) (l_e_st_eq_landau_n_rt_rp_r_absn m nm) (l_e_st_eq_landau_n_rt_rp_r_absn n nn)) (l_e_st_eq_landau_n_rt_rp_r_satz180a m n) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_pl m n)) (l_e_st_eq_landau_n_rt_rp_r_absn (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_negpl m n nm nn)))))))))))). Time Defined. (* constant 6565 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_negpl m n nm nn)))))))))). Time Defined. (* constant 6566 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t31 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t27 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t29 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t27 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6567 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t32 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t24 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t25 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t27 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t26 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t28 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t31 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6568 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t33 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_satz166b m nm) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) nm)))))))))). Time Defined. (* constant 6569 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t34 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_satz166b n nn) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) nn)))))))))). Time Defined. (* constant 6570 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t35 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_satz166b (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_negpl m n nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_negpl m n nm nn))))))))))). Time Defined. (* constant 6571 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t36 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t33 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t34 x m n mi ni o na nm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t33 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t34 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_negexp x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) nm) (l_e_st_eq_landau_n_rt_rp_r_c_negexp x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) nn)))))))))). Time Defined. (* constant 6572 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t37 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn))) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_satz221d (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t33 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t34 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6573 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t38 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t33 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t34 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t37 x m n mi ni o na nm nn))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_satz247 l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t33 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t34 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6574 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t39 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t37 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t35 x m n mi ni o na nm nn))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_isov12 (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_satz222 l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t32 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t37 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t35 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6575 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t40 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t35 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t35 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_negexp x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_negpl m n nm nn))))))))))). Time Defined. (* constant 6576 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t41 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t33 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t34 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_am x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t17 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t22 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t23 x m n mi ni o na nm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t37 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t30 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t35 x m n mi ni o na nm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t36 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t38 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t39 x m n mi ni o na nm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t40 x m n mi ni o na nm nn)))))))))). Time Defined. (* constant 6577 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) nn))))))))). Time Defined. (* constant 6578 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn) (l_e_st_eq_landau_n_rt_rp_r_satz166b n nn) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) nn)))))))))). Time Defined. (* constant 6579 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t44 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ists2 (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_negexp x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) nn)))))))))). Time Defined. (* constant 6580 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t45 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_satz244a (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)))))))))). Time Defined. (* constant 6581 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t46 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_isov1 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_satz222 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)))))))))). Time Defined. (* constant 6582 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t47 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t44 x m n mi ni o na pm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t45 x m n mi ni o na pm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t46 x m n mi ni o na pm nn)))))))))). Time Defined. (* constant 6583 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t48 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_satz182d m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) casea)))))))))). Time Defined. (* constant 6584 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t49 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_satz166e n (l_e_st_eq_landau_n_rt_rp_r_nnot0 n nn))))))))))). Time Defined. (* constant 6585 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t50 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t49 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t48 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6586 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t51 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t50 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6587 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_intmn m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o))))))))))). Time Defined. (* constant 6588 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t51 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6589 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t51 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6590 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t55 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t51 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6591 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t56 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6592 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t57 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t56 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t55 x m n mi ni o na pm nn casea)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t14 x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t51 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t50 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6593 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t58 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) m)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_satz187a m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))). Time Defined. (* constant 6594 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t59 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t56 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t55 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) m (l_e_st_eq_landau_n_rt_rp_r_c_9294_t58 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t56 x m n mi ni o na pm nn casea) mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t55 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o))))))))))). Time Defined. (* constant 6595 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t60 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t56 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t55 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t57 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t59 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6596 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t61 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn))))))))))). Time Defined. (* constant 6597 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t62 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_isp1 l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_nis t l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t61 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6598 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t63 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t62 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_isov12 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t60 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t61 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t62 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn))))))))))). Time Defined. (* constant 6599 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t64 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t62 x m n mi ni o na pm nn casea)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t62 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t47 x m n mi ni o na pm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t63 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6600 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t65 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t62 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_satz229h (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t62 x m n mi ni o na pm nn casea) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea)))))))))))))). Time Defined. (* constant 6601 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t66 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pl m n))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_m0 n)) (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_ismn2 (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_m0 n) m (l_e_st_eq_landau_n_rt_rp_r_absn n nn)) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 n)) n m (l_e_st_eq_landau_n_rt_rp_r_satz177 n)))))))))))). Time Defined. (* constant 6602 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t67 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t66 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o))))))))))). Time Defined. (* constant 6603 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t68 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casea:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t53 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t62 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t52 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t54 x m n mi ni o na pm nn casea)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t64 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t65 x m n mi ni o na pm nn casea) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t67 x m n mi ni o na pm nn casea))))))))))). Time Defined. (* constant 6604 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t69 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) caseb mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn))))))))))). Time Defined. (* constant 6605 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t70 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)) l_e_st_eq_landau_n_rt_rp_r_c_1c)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_satz251a (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t69 x m n mi ni o na pm nn caseb))))))))))). Time Defined. (* constant 6606 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t71 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl m n) l_e_st_eq_landau_n_rt_rp_r_0)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_m0 n)) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ispl2 n (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 n)) m (l_e_st_eq_landau_n_rt_rp_r_satz177a n)) (l_e_st_eq_landau_n_rt_rp_r_ismn2 (l_e_st_eq_landau_n_rt_rp_r_m0 n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_m0 n) (l_e_st_eq_landau_n_rt_rp_r_absn n nn))) (l_e_st_eq_landau_n_rt_rp_r_satz182e m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) caseb))))))))))). Time Defined. (* constant 6607 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t72 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_0exp x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t71 x m n mi ni o na pm nn caseb)))))))))))). Time Defined. (* constant 6608 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t73 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (caseb:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t47 x m n mi ni o na pm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t70 x m n mi ni o na pm nn caseb) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t72 x m n mi ni o na pm nn caseb))))))))))). Time Defined. (* constant 6609 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t74 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_satz182d (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m (l_e_st_eq_landau_n_rt_rp_r_lemma2 m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) casec))))))))))). Time Defined. (* constant 6610 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t75 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_andi (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)) pm (l_e_st_eq_landau_n_rt_rp_r_c_9294_t74 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6611 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t76 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t75 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6612 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_intmn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) m mi)))))))))). Time Defined. (* constant 6613 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t76 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6614 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t76 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6615 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t80 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m))))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t76 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6616 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t81 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_pl m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_intpl m mi (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6617 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t81a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t81 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t80 x m n mi ni o na pm nn casec)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t14 x m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t76 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t75 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6618 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t82 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_satz187a (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)))))))))). Time Defined. (* constant 6619 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t83 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t81 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t80 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_pl m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t82 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t81 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t80 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn))))))))))). Time Defined. (* constant 6620 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t84 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t81 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t80 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t81a x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t83 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6621 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t85 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t15 x m n mi ni o na))))))))))). Time Defined. (* constant 6622 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t86 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t15 x m n mi ni o na))))))))))). Time Defined. (* constant 6623 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t87 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_satz221d (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t85 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t86 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6624 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t88 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_satz222 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)))))))))))). Time Defined. (* constant 6625 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t89 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o))))))))))). Time Defined. (* constant 6626 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t90 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t88 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t89 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6627 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t91 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t87 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_isov12 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t90 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t84 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t87 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn))))))))))). Time Defined. (* constant 6628 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t92 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t87 x m n mi ni o na pm nn casec)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t87 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t18 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t42 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t43 x m n mi ni o na pm nn)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t47 x m n mi ni o na pm nn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t91 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6629 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t93 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t87 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t86 x m n mi ni o na pm nn casec)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_satz246a l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t86 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t85 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6630 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t94 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_satz182f m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) casec)))))))))). Time Defined. (* constant 6631 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t94a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pl m n))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_m0 n)) (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_ismn2 (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_m0 n) m (l_e_st_eq_landau_n_rt_rp_r_absn n nn)) (l_e_st_eq_landau_n_rt_rp_r_ispl2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_m0 n)) n m (l_e_st_eq_landau_n_rt_rp_r_satz177 n)))))))))))). Time Defined. (* constant 6632 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t95 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_satz181a (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_abs (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_absn (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t94 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_isabs (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t94a x m n mi ni o na pm nn casec)))))))))))). Time Defined. (* constant 6633 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t96 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_pl m n))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_real (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_neg t) (l_e_st_eq_landau_n_rt_rp_r_mn m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t94 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t94a x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6634 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw3 x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t96 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6635 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t98 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_lemmapw1 x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_satz166b (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t96 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_lemmapw2 x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t96 x m n mi ni o na pm nn casec)))))))))))). Time Defined. (* constant 6636 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t99 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t95 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6637 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t100 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t86 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t98 x m n mi ni o na pm nn casec)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_isov2 (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_9294_t99 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t86 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t98 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6638 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t101 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t98 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t98 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_negexp x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t96 x m n mi ni o na pm nn casec)))))))))))). Time Defined. (* constant 6639 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t102 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), (forall (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => (fun (casec:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) l_e_st_eq_landau_n_rt_rp_r_c_1c) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t78 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec))) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t87 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o) m) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t77 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t79 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t86 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_c_9294_ap x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t19 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t97 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t98 x m n mi ni o na pm nn casec)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t92 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t93 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t100 x m n mi ni o na pm nn casec) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t101 x m n mi ni o na pm nn casec))))))))))). Time Defined. (* constant 6640 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t103 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), (forall (nn:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => (fun (nn:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_or3app (l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_satz167a m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t73 x m n mi ni o na pm nn t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_more m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t68 x m n mi ni o na pm nn t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_less m (l_e_st_eq_landau_n_rt_rp_r_c_9294_an x m n mi ni o)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t102 x m n mi ni o na pm nn t)))))))))). Time Defined. (* constant 6641 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t104 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos n) (l_e_st_eq_landau_n_rt_rp_r_pos m))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos n) (l_e_st_eq_landau_n_rt_rp_r_pos m)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t15 x m n mi ni o na)))))))). Time Defined. (* constant 6642 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t104a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos n) (l_e_st_eq_landau_n_rt_rp_r_pos m))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_and_th5 (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n) na))))))). Time Defined. (* constant 6643 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x n m ni mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t104 x m n mi ni o na)))))))). Time Defined. (* constant 6644 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x n m ni mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t104 x m n mi ni o na)))))))). Time Defined. (* constant 6645 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t107 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl n m))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x n m ni mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t104 x m n mi ni o na)))))))). Time Defined. (* constant 6646 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t108 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x n (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na))))))))). Time Defined. (* constant 6647 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t109 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_comts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na))))))))). Time Defined. (* constant 6648 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t110 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (qn:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl n m) (l_e_st_eq_landau_n_rt_rp_r_intpl n ni m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t107 x m n mi ni o na))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (qn:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t103 x n m ni mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t104 x m n mi ni o na) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t104a x m n mi ni o na) qn nm))))))))). Time Defined. (* constant 6649 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t111 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl n m) (l_e_st_eq_landau_n_rt_rp_r_intpl n ni m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t107 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_pl n m) (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_compl n m) (l_e_st_eq_landau_n_rt_rp_r_intpl n ni m mi) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t107 x m n mi ni o na) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)))))))). Time Defined. (* constant 6650 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t112 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), (forall (qn:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => (fun (qn:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl n m) (l_e_st_eq_landau_n_rt_rp_r_intpl n ni m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t107 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t108 x m n mi ni o na) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t109 x m n mi ni o na) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t110 x m n mi ni o na nm qn) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t111 x m n mi ni o na)))))))))). Time Defined. (* constant 6651 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t113 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_ists1 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_0exp x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) i))))))))). Time Defined. (* constant 6652 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t114 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_satz222b (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)))))))))). Time Defined. (* constant 6653 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t115 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is n (l_e_st_eq_landau_n_rt_rp_r_pl m n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl m n) n (l_e_st_eq_landau_n_rt_rp_r_pl01 m n i))))))))). Time Defined. (* constant 6654 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t116 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x n (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t115 x m n mi ni o na i) ni (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o))))))))). Time Defined. (* constant 6655 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t117 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t113 x m n mi ni o na i) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t114 x m n mi ni o na i) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t116 x m n mi ni o na i))))))))). Time Defined. (* constant 6656 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t118 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl n m) (l_e_st_eq_landau_n_rt_rp_r_intpl n ni m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t107 x m n mi ni o na)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t117 x n m ni mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t104 x m n mi ni o na) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t104a x m n mi ni o na) i)))))))). Time Defined. (* constant 6657 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t119 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9294_t105 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9294_t106 x m n mi ni o na))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl n m) (l_e_st_eq_landau_n_rt_rp_r_intpl n ni m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t107 x m n mi ni o na)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t108 x m n mi ni o na) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t109 x m n mi ni o na) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t118 x m n mi ni o na i) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t111 x m n mi ni o na))))))))). Time Defined. (* constant 6658 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t120 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), l_not (l_e_st_eq_landau_n_rt_rp_r_pos n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_ore2 (l_not (l_e_st_eq_landau_n_rt_rp_r_pos m)) (l_not (l_e_st_eq_landau_n_rt_rp_r_pos n)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t16 x m n mi ni o na) (l_weli (l_e_st_eq_landau_n_rt_rp_r_pos m) pm))))))))). Time Defined. (* constant 6659 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t121 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (pm:l_e_st_eq_landau_n_rt_rp_r_pos m), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (pm:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_rapp n (l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o) (l_imp_th2 (l_e_st_eq_landau_n_rt_rp_r_pos n) (l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9294_t120 x m n mi ni o na pm)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t119 x m n mi ni o na t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t103 x m n mi ni o na pm t))))))))). Time Defined. (* constant 6660 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t122 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (nm:l_e_st_eq_landau_n_rt_rp_r_neg m), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (nm:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_rapp n (l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t112 x m n mi ni o na nm t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t119 x m n mi ni o na t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t41 x m n mi ni o na nm t))))))))). Time Defined. (* constant 6661 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9294_t123 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (na:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_rapp m (l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos m) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t121 x m n mi ni o na t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is m l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t117 x m n mi ni o na t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg m) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t122 x m n mi ni o na t)))))))). Time Defined. (* constant 6662 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz294 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl m n) (l_e_st_eq_landau_n_rt_rp_r_intpl m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_imp_th1 (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) (l_e_st_eq_landau_n_rt_rp_r_c_9294_prop1 x m n mi ni o) (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t14 x m n mi ni o t) (fun (t:l_not (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_9294_t123 x m n mi ni o t))))))). Time Defined. (* constant 6663 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma295a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m) o)))))). Time Defined. (* constant 6664 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma295b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n) o)))))). Time Defined. (* constant 6665 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma295c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m n)) o)))))). Time Defined. (* constant 6666 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m n)) (l_e_st_eq_landau_n_rt_rp_r_pos n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m n)) (l_e_st_eq_landau_n_rt_rp_r_pos n)) o)))))). Time Defined. (* constant 6667 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_mn m n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_intmn m mi n ni)))))). Time Defined. (* constant 6668 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn m n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x (l_e_st_eq_landau_n_rt_rp_r_mn m n) n (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) ni (l_e_st_eq_landau_n_rt_rp_r_c_9295_t1 x m n mi ni o))))))). Time Defined. (* constant 6669 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x (l_e_st_eq_landau_n_rt_rp_r_mn m n) n (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) ni (l_e_st_eq_landau_n_rt_rp_r_c_9295_t1 x m n mi ni o))))))). Time Defined. (* constant 6670 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn m n) n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x (l_e_st_eq_landau_n_rt_rp_r_mn m n) n (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) ni (l_e_st_eq_landau_n_rt_rp_r_c_9295_t1 x m n mi ni o))))))). Time Defined. (* constant 6671 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma295c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t3 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9295_t4 x m n mi ni o))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma295c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t3 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9295_t4 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma295c x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t3 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x n (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t4 x m n mi ni o)))))))). Time Defined. (* constant 6672 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t3 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9295_t4 x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn m n) n) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) n ni) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t5 x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz294 x (l_e_st_eq_landau_n_rt_rp_r_mn m n) n (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) ni (l_e_st_eq_landau_n_rt_rp_r_c_9295_t1 x m n mi ni o))))))). Time Defined. (* constant 6673 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn m n) n) m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_plmn m n)))))). Time Defined. (* constant 6674 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn m n) n) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) n ni) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t5 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma295a x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn m n) n) m (l_e_st_eq_landau_n_rt_rp_r_c_9295_t8 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) n ni) mi (l_e_st_eq_landau_n_rt_rp_r_c_9295_t5 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma295a x m n mi ni o))))))). Time Defined. (* constant 6675 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma295c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma295a x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma295c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t3 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_9295_t4 x m n mi ni o))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_mn m n) n) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) n ni) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t5 x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma295a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t6 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t7 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t9 x m n mi ni o))))))). Time Defined. (* constant 6676 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9295_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o) o)))))). Time Defined. (* constant 6677 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz295 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma295a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_satz290 x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o) o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_intmn m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma295c x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz229k (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma295a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn m n) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t2 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_lemma295c x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t11 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9295_t10 x m n mi ni o))))))). Time Defined. (* constant 6678 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma296 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m) n)))). Time Defined. (* constant 6679 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_intrl l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_intrli0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 6680 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma295a x l_e_st_eq_landau_n_rt_rp_r_0 m (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) mi n)))). Time Defined. (* constant 6681 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma295b x l_e_st_eq_landau_n_rt_rp_r_0 m (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) mi n)))). Time Defined. (* constant 6682 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma295c x l_e_st_eq_landau_n_rt_rp_r_0 m (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) mi n)))). Time Defined. (* constant 6683 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n) n)))). Time Defined. (* constant 6684 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 x m mi n)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 x m mi n) n)))). Time Defined. (* constant 6685 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t2 x m mi n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t2 x m mi n)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_0exp x l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t2 x m mi n) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))))). Time Defined. (* constant 6686 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 x m mi n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 x m mi n))))). Time Defined. (* constant 6687 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t5 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t2 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t6 x m mi n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_isov12 l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t2 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t7 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t8 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t5 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t6 x m mi n))))). Time Defined. (* constant 6688 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t2 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t6 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 m) (l_e_st_eq_landau_n_rt_rp_r_intmn l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t4 x m mi n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz295 x l_e_st_eq_landau_n_rt_rp_r_0 m (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) mi n)))). Time Defined. (* constant 6689 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 m) (l_e_st_eq_landau_n_rt_rp_r_m0 m))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 6690 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 m)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_intm0 m mi) n)))). Time Defined. (* constant 6691 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 m) (l_e_st_eq_landau_n_rt_rp_r_intmn l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t4 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_intm0 m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t12 x m mi n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 m) (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t11 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_intmn l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) m mi) (l_e_st_eq_landau_n_rt_rp_r_intm0 m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t4 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t12 x m mi n))))). Time Defined. (* constant 6692 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9296_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t5 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_intm0 m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t12 x m mi n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t5 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pw x l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t2 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9296_t3 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t6 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 m) (l_e_st_eq_landau_n_rt_rp_r_intmn l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_9296_t1 x m mi n) m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t4 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_intm0 m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t12 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t9 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t10 x m mi n) (l_e_st_eq_landau_n_rt_rp_r_c_9296_t13 x m mi n))))). Time Defined. (* constant 6693 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz296 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n)) (l_e_st_eq_landau_n_rt_rp_r_c_satz290 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x m mi n) n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_intm0 m mi) (l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x (l_e_st_eq_landau_n_rt_rp_r_m0 m) (l_e_st_eq_landau_n_rt_rp_r_intm0 m mi) n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_9296_t14 x m mi n)))). Time Defined. (* constant 6694 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) p))))))). Time Defined. (* constant 6695 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_pos n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_ande2 (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n) a))))))). Time Defined. (* constant 6696 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t3 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (a:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_postspp m n (l_ande1 (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n) a) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t2 x m n mi ni o a)))))))). Time Defined. (* constant 6697 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma297a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_or_th9 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos n) o (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t1 x m n mi ni o t) (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t2 x m n mi ni o t))))))). Time Defined. (* constant 6698 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma297b : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_or_th8 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m n)) o (fun (t:l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n)) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t3 x m n mi ni o t))))))). Time Defined. (* constant 6699 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t4 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_pos m))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ore2 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m) o (l_weli (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) i)))))). Time Defined. (* constant 6700 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_nat))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_ntofrl m (l_e_st_eq_landau_n_rt_rp_r_posintnatrl m (l_e_st_eq_landau_n_rt_rp_r_c_9297_t4 x m mi o i) mi)))))). Time Defined. (* constant 6701 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t5 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i)) => x))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_posexp x m mi o (l_e_st_eq_landau_n_rt_rp_r_c_9297_t4 x m mi o i)))))). Time Defined. (* constant 6702 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t6 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i)) => x)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz289b (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i)) => x) (l_e_st_eq_landau_n_xout (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i)) i))))). Time Defined. (* constant 6703 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t7 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) (l_e_st_eq_landau_n_rt_rp_r_c_prod (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i) (fun (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_9297_m1 x m mi o i)) => x)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_9297_t5 x m mi o i) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t6 x m mi o i)))))). Time Defined. (* constant 6704 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_pw0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi o) l_e_st_eq_landau_n_rt_rp_r_c_0c))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m)) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t7 x m mi o i))))). Time Defined. (* constant 6705 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_ts m n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => l_e_st_eq_landau_n_rt_rp_r_intts m mi n ni))))). Time Defined. (* constant 6706 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_prop1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), Prop)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)))))))). Time Defined. (* constant 6707 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t9 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_pw0 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) i))))))). Time Defined. (* constant 6708 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t10 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_pw0 (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t9 x m n mi ni o i)))))))). Time Defined. (* constant 6709 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t11 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_0c))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_pw0 x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o) i))))))). Time Defined. (* constant 6710 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t12 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop1 x m n mi ni o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_0c (l_e_st_eq_landau_n_rt_rp_r_c_9297_t10 x m n mi ni o i) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t11 x m n mi ni o i)))))))). Time Defined. (* constant 6711 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t13 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m) p)))). Time Defined. (* constant 6712 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_cx)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t13 x m mi p))))). Time Defined. (* constant 6713 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_nr : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_real))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_rlofnt n))))). Time Defined. (* constant 6714 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_natintrl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_natrli n)))))). Time Defined. (* constant 6715 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => l_ori2 (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_natpos (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_natrli n))))))). Time Defined. (* constant 6716 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n))) p))))). Time Defined. (* constant 6717 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_intts m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n)))))). Time Defined. (* constant 6718 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), Prop))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p n))))))). Time Defined. (* constant 6719 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t18 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_ispw1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p)) l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p)))))). Time Defined. (* constant 6720 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t19 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_satz291 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p))))). Time Defined. (* constant 6721 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t20 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p l_e_st_eq_landau_n_1)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x m (l_e_st_eq_landau_n_rt_rp_r_ts m l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_satz195a m) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t13 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p l_e_st_eq_landau_n_1))))). Time Defined. (* constant 6722 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t21 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p l_e_st_eq_landau_n_1)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p l_e_st_eq_landau_n_1)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t18 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t19 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t20 x m mi p))))). Time Defined. (* constant 6723 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_nat)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_pl n l_e_st_eq_landau_n_1)))))). Time Defined. (* constant 6724 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t22 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t13 x m mi p) p)))))). Time Defined. (* constant 6725 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t23 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_1rl)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t22 x m mi p n p2))))))). Time Defined. (* constant 6726 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294a (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t23 x m mi p n p2))))))). Time Defined. (* constant 6727 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos l_e_st_eq_landau_n_rt_rp_r_1rl))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294b (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t23 x m mi p n p2))))))). Time Defined. (* constant 6728 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t26 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294c (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t23 x m mi p n p2))))))). Time Defined. (* constant 6729 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t27 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t26 x m mi p n p2)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_satzr155a n l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t26 x m mi p n p2))))))). Time Defined. (* constant 6730 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t27a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t26 x m mi p n p2)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_satz294 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t23 x m mi p n p2))))))). Time Defined. (* constant 6731 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t28 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t26 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t27 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t27a x m mi p n p2))))))). Time Defined. (* constant 6732 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t29 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n))) (l_e_st_eq_landau_n_rt_rp_r_pos m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n))) (l_e_st_eq_landau_n_rt_rp_r_pos m)) p)))))). Time Defined. (* constant 6733 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t30 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t29 x m mi p n p2))))))). Time Defined. (* constant 6734 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t31 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294b x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t29 x m mi p n p2))))))). Time Defined. (* constant 6735 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t32 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294c x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t29 x m mi p n p2))))))). Time Defined. (* constant 6736 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t33 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t30 x m mi p n p2)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t30 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p n)) p2 (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t30 x m mi p n p2)))))))). Time Defined. (* constant 6737 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t34 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t31 x m mi p n p2)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t31 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p)) l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_c_lemma291 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t19 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t13 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t31 x m mi p n p2)))))))). Time Defined. (* constant 6738 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t35 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t30 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t31 x m mi p n p2))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t30 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t31 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t33 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t34 x m mi p n p2))))))). Time Defined. (* constant 6739 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t36 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t30 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t31 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t32 x m mi p n p2)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_satz294 x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t29 x m mi p n p2))))))). Time Defined. (* constant 6740 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t37 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m) (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_ts m l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_ispl2 m (l_e_st_eq_landau_n_rt_rp_r_ts m l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_satz195a m)) (l_e_st_eq_landau_n_rt_rp_r_distpt2 m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) m (l_e_st_eq_landau_n_rt_rp_r_satzr155b n l_e_st_eq_landau_n_1)))))))). Time Defined. (* constant 6741 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t38 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t32 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m) (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t37 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t32 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)))))))). Time Defined. (* constant 6742 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t39 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t24 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_intrl1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_t25 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t30 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t31 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) m) (l_e_st_eq_landau_n_rt_rp_r_intpl (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p n)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p n) m mi) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t32 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_nr x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t28 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t35 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t36 x m mi p n p2) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t38 x m mi p n p2))))))). Time Defined. (* constant 6743 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t40 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), (forall (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p (l_e_st_eq_landau_n_suc n))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p2:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p t) (l_e_st_eq_landau_n_rt_rp_r_c_9297_n1 x m mi p n p2) (l_e_st_eq_landau_n_suc n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t39 x m mi p n p2) (l_e_st_eq_landau_n_satz4a n))))))). Time Defined. (* constant 6744 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t41 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p n))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p t) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t21 x m mi p) (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_9297_prop2 x m mi p t) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t40 x m mi p t u)) n))))). Time Defined. (* constant 6745 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t42 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_natrl n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_posintnatrl n q ni)))))))). Time Defined. (* constant 6746 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_nat)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_ntofrl n (l_e_st_eq_landau_n_rt_rp_r_c_9297_t42 x m n mi ni o p q))))))))). Time Defined. (* constant 6747 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t43 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_is n (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_isrlnt1 n (l_e_st_eq_landau_n_rt_rp_r_c_9297_t42 x m n mi ni o p q))))))))). Time Defined. (* constant 6748 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t44 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) n)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_isrlnt2 n (l_e_st_eq_landau_n_rt_rp_r_c_9297_t42 x m n mi ni o p q))))))))). Time Defined. (* constant 6749 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o))))))). Time Defined. (* constant 6750 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t44a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x m m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real m) mi mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t13 x m mi p))))))))). Time Defined. (* constant 6751 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t45 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw12 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) n (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t44a x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t43 x m n mi ni o p q) ni (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)))))))))). Time Defined. (* constant 6752 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t46 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t41 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))))))))). Time Defined. (* constant 6753 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t47 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_ts m n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) n m (l_e_st_eq_landau_n_rt_rp_r_c_9297_t44 x m n mi ni o p q))))))))). Time Defined. (* constant 6754 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t48 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t47 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o))))))))). Time Defined. (* constant 6755 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t49 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_pos n), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop1 x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p0 x m mi p) (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t14 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t15 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_rlofnt (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t17 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t16 x m mi p (l_e_st_eq_landau_n_rt_rp_r_c_9297_n0 x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t45 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t46 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t48 x m n mi ni o p q))))))))). Time Defined. (* constant 6756 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t50 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_0exp (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o) i)))))))). Time Defined. (* constant 6757 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t51 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts m n) l_e_st_eq_landau_n_rt_rp_r_0)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_ts02 m n i)))))))). Time Defined. (* constant 6758 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t52 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_0exp x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t51 x m n mi ni o p i))))))))). Time Defined. (* constant 6759 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t53 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop1 x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_9297_t50 x m n mi ni o p i) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t52 x m n mi ni o p i))))))))). Time Defined. (* constant 6760 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_an : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_real)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_abs n)))))))). Time Defined. (* constant 6761 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_intabs n ni)))))))). Time Defined. (* constant 6762 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t55 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_ori1 (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) p)))))))). Time Defined. (* constant 6763 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_cx)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t55 x m n mi ni o p q))))))))). Time Defined. (* constant 6764 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t56 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_satz166e n (l_e_st_eq_landau_n_rt_rp_r_nnot0 n q))))))))). Time Defined. (* constant 6765 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t56a : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos m))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t55 x m n mi ni o p q))))))))). Time Defined. (* constant 6766 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t57 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t55 x m n mi ni o p q))))))))). Time Defined. (* constant 6767 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t58 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t55 x m n mi ni o p q))))))))). Time Defined. (* constant 6768 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t59 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t57 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t58 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t49 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t55 x m n mi ni o p q) p (l_e_st_eq_landau_n_rt_rp_r_c_9297_t56 x m n mi ni o p q))))))))). Time Defined. (* constant 6769 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t60 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_is n (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_satz177c (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) n (l_e_st_eq_landau_n_rt_rp_r_absn n q))))))))). Time Defined. (* constant 6770 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t61 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_intm0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q))))))))). Time Defined. (* constant 6771 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t62 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x m mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t56a x m n mi ni o p q) p)))))))). Time Defined. (* constant 6772 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma296 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t62 x m n mi ni o p q))))))))). Time Defined. (* constant 6773 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t64 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t62 x m n mi ni o p q))))))))). Time Defined. (* constant 6774 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t65 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma296 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t61 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t62 x m n mi ni o p q))))))))). Time Defined. (* constant 6775 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t66 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t64 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t61 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t65 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_satz296 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t62 x m n mi ni o p q))))))))). Time Defined. (* constant 6776 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t67 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x m (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t56a x m n mi ni o p q))))))))). Time Defined. (* constant 6777 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t68 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t61 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t65 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw12 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) n (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t67 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t60 x m n mi ni o p q) ni (l_e_st_eq_landau_n_rt_rp_r_c_9297_t61 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t65 x m n mi ni o p q))))))))). Time Defined. (* constant 6778 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t69 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t64 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t64 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t61 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t65 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t68 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t66 x m n mi ni o p q))))))))). Time Defined. (* constant 6779 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t70 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_ts m n))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_satz197f m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_ists2 (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) n m (l_e_symis l_e_st_eq_landau_n_rt_rp_r_real n (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t60 x m n mi ni o p q))))))))))). Time Defined. (* constant 6780 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t71 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_intrl (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_intm0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)))))))))). Time Defined. (* constant 6781 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) p)))))))). Time Defined. (* constant 6782 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t73 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q)) l_e_st_eq_landau_n_rt_rp_r_c_0c)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_satz290 x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q) p)))))))). Time Defined. (* constant 6783 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t74 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_pos (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)))))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_lemma296 x (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t71 x m n mi ni o p q) p)))))))). Time Defined. (* constant 6784 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t75 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t73 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t71 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t74 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_satz296 x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) p)))))))). Time Defined. (* constant 6785 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t76 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t71 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t74 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw2 x (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t70 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t71 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t74 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o))))))))). Time Defined. (* constant 6786 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t77 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t57 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw1 (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t57 x m n mi ni o p q))))))))). Time Defined. (* constant 6787 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t78 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t58 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_ispw1 x x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_refis l_e_st_eq_landau_n_rt_rp_r_c_cx x) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t58 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q))))))))). Time Defined. (* constant 6788 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t79 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t57 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t58 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t77 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t59 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t78 x m n mi ni o p q))))))))). Time Defined. (* constant 6789 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t80 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t64 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t73 x m n mi ni o p q)))))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_isov2 (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q)) l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_9297_t79 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t64 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t73 x m n mi ni o p q))))))))). Time Defined. (* constant 6790 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t81 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), (forall (q:l_e_st_eq_landau_n_rt_rp_r_neg n), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop1 x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => (fun (q:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_tr4is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1 x m n mi ni o) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_9297_p1t55 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t63 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t64 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_ov l_e_st_eq_landau_n_rt_rp_r_c_1c (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q) mi (l_e_st_eq_landau_n_rt_rp_r_c_9297_t54 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t72 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t73 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts m (l_e_st_eq_landau_n_rt_rp_r_c_9297_an x m n mi ni o p q))) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t71 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t74 x m n mi ni o p q)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t8 x m n mi ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t69 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t80 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t75 x m n mi ni o p q) (l_e_st_eq_landau_n_rt_rp_r_c_9297_t76 x m n mi ni o p q))))))))). Time Defined. (* constant 6791 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_9297_t82 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_c_9297_prop1 x m n mi ni o))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_rapp n (l_e_st_eq_landau_n_rt_rp_r_c_9297_prop1 x m n mi ni o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos n) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t49 x m n mi ni o p t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is n l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t53 x m n mi ni o p t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg n) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t81 x m n mi ni o p t)))))))). Time Defined. (* constant 6792 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz297 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), (forall (m:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_real), (forall (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m), (forall (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n), (forall (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pw (l_e_st_eq_landau_n_rt_rp_r_c_pw x m mi (l_e_st_eq_landau_n_rt_rp_r_c_lemma294a x m n mi ni o)) n ni (l_e_st_eq_landau_n_rt_rp_r_c_lemma297a x m n mi ni o)) (l_e_st_eq_landau_n_rt_rp_r_c_pw x (l_e_st_eq_landau_n_rt_rp_r_ts m n) (l_e_st_eq_landau_n_rt_rp_r_intts m mi n ni) (l_e_st_eq_landau_n_rt_rp_r_c_lemma297b x m n mi ni o)))))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => (fun (m:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (mi:l_e_st_eq_landau_n_rt_rp_r_intrl m) => (fun (ni:l_e_st_eq_landau_n_rt_rp_r_intrl n) => (fun (o:l_or (l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_and (l_e_st_eq_landau_n_rt_rp_r_pos m) (l_e_st_eq_landau_n_rt_rp_r_pos n))) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_c_9297_prop1 x m n mi ni o) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t12 x m n mi ni o t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_nis x l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_9297_t82 x m n mi ni o t))))))). Time Defined. (* constant 6793 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_plis12a r l_e_st_eq_landau_n_rt_rp_r_0 s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl r s) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 6794 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t1 r s))). Time Defined. (* constant 6795 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_10298_t1 r s)). Time Defined. (* constant 6796 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_mnis12a r l_e_st_eq_landau_n_rt_rp_r_0 s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 (l_e_st_eq_landau_n_rt_rp_r_mn l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_mn r s) (l_e_st_eq_landau_n_rt_rp_r_pl02 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_satz176b l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))))). Time Defined. (* constant 6797 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298c : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t2 r s))). Time Defined. (* constant 6798 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298d : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_mn (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_10298_t2 r s)). Time Defined. (* constant 6799 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_satz176b (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 6800 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s)) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s)) (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 s (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ts02 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 6801 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s))) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_tsis12a r l_e_st_eq_landau_n_rt_rp_r_0 s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts r s) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 s)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_10298_t3 r s) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t4 r s)))). Time Defined. (* constant 6802 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t5 r s))). Time Defined. (* constant 6803 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts r s) l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_10298_t5 r s)). Time Defined. (* constant 6804 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_0c), l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real s (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_re l_e_st_eq_landau_n_rt_rp_r_c_0c) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_c_isre s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_iscere (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_0c i) (l_e_st_eq_landau_n_rt_rp_r_c_reis l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 6805 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_lemma298 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_nis (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_0c))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_imp_th3 (l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_0c) (l_e_st_eq_landau_n_rt_rp_r_is s l_e_st_eq_landau_n_rt_rp_r_0) n (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_c_0c) => l_e_st_eq_landau_n_rt_rp_r_c_10298_t6 r s n t)))). Time Defined. (* constant 6806 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ov r s n) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ov r s n) l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t5 s (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx1 (l_e_st_eq_landau_n_rt_rp_r_ts s (l_e_st_eq_landau_n_rt_rp_r_ov r s n)) r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_satz204c r s n))))). Time Defined. (* constant 6807 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298g : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ov r s n) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_lemma298 r s n))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_satz229g (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ov r s n) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_lemma298 r s n) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t7 r s n)))). Time Defined. (* constant 6808 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298h : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ov (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_lemma298 r s n)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ov r s n) l_e_st_eq_landau_n_rt_rp_r_0)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_nis s l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_satz229h (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_ov r s n) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_lemma298 r s n) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t7 r s n)))). Time Defined. (* constant 6809 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_m0isa r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_isrecx2 (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_m0 r) (l_e_st_eq_landau_n_rt_rp_r_satz176b l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 6810 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298j : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t8 r)). Time Defined. (* constant 6811 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298k : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_m0 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_m0 r) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_10298_t8 r). Time Defined. (* constant 6812 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod2 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts r r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod2 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_imis r l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ists12 (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) r (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) r (l_e_st_eq_landau_n_rt_rp_r_c_reis r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_reis r l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 6813 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_ar : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_real). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_abs r). Time Defined. (* constant 6814 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t10 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (p:l_e_st_eq_landau_n_rt_rp_r_pos r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_satz196a r r p p)). Time Defined. (* constant 6815 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t11 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_tris2 l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_ts01 r r i) (l_e_st_eq_landau_n_rt_rp_r_ts01 (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_abs0 r i)))). Time Defined. (* constant 6816 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t12 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (n:l_e_st_eq_landau_n_rt_rp_r_neg r), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (n:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_satz196b r r n n)). Time Defined. (* constant 6817 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t13 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_rapp r (l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_pos r) => l_e_st_eq_landau_n_rt_rp_r_c_10298_t10 r t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_c_10298_t11 r t) (fun (t:l_e_st_eq_landau_n_rt_rp_r_neg r) => l_e_st_eq_landau_n_rt_rp_r_c_10298_t12 r t)). Time Defined. (* constant 6818 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t14 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod2 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod2 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts r r) (l_e_st_eq_landau_n_rt_rp_r_ts (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r)) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t9 r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t13 r)). Time Defined. (* constant 6819 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10298_t15 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_imp_th1 (l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) (l_not (l_e_st_eq_landau_n_rt_rp_r_neg (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r))) (fun (t:l_e_st_eq_landau_n_rt_rp_r_is r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_0notn (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_abs0 r t)) (fun (t:l_e_st_eq_landau_n_rt_rp_r_nis r l_e_st_eq_landau_n_rt_rp_r_0) => l_e_st_eq_landau_n_rt_rp_r_pnotn (l_e_st_eq_landau_n_rt_rp_r_c_10298_ar r) (l_e_st_eq_landau_n_rt_rp_r_satz166e r t))). Time Defined. (* constant 6820 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298l : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_abs r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_thsqrt3 (l_e_st_eq_landau_n_rt_rp_r_c_mod2 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_c_lemma5 (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t15 r) (l_e_st_eq_landau_n_rt_rp_r_c_10298_t14 r)). Time Defined. (* constant 6821 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz298m : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_c_mod (l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_abs r) (l_e_st_eq_landau_n_rt_rp_r_c_satz298l r)). Time Defined. (* constant 6822 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_cofrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_complex). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_pli r l_e_st_eq_landau_n_rt_rp_r_0). Time Defined. (* constant 6823 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isrlic : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s)), l_e_st_eq_landau_n_rt_rp_r_is r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s)) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s)) s (l_e_st_eq_landau_n_rt_rp_r_c_isre r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_c_iscere (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) i) (l_e_st_eq_landau_n_rt_rp_r_c_reis s l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 6824 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isrlec : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is r s), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is r s) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx1 r s l_e_st_eq_landau_n_rt_rp_r_0 i))). Time Defined. (* constant 6825 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v10_t1 : l_e_injective l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) (l_e_st_eq_landau_n_rt_rp_r_c_cofrl u)) => l_e_st_eq_landau_n_rt_rp_r_c_isrlic t u v))). Time Defined. (* constant 6826 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_realc : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), Prop). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_image l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) x). Time Defined. (* constant 6827 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_reali : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_realc (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_imagei l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) r). Time Defined. (* constant 6828 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_rlofc : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x), l_e_st_eq_landau_n_rt_rp_r_real)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x) => l_e_soft l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t1 x rx)). Time Defined. (* constant 6829 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscirl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (ry:l_e_st_eq_landau_n_rt_rp_r_c_realc y), (forall (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rlofc x rx) (l_e_st_eq_landau_n_rt_rp_r_c_rlofc y ry)), l_e_st_eq_landau_n_rt_rp_r_c_is x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (ry:l_e_st_eq_landau_n_rt_rp_r_c_realc y) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rlofc x rx) (l_e_st_eq_landau_n_rt_rp_r_c_rlofc y ry)) => l_e_isinve l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t1 x rx y ry i))))). Time Defined. (* constant 6830 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscerl : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (ry:l_e_st_eq_landau_n_rt_rp_r_c_realc y), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rlofc x rx) (l_e_st_eq_landau_n_rt_rp_r_c_rlofc y ry)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (ry:l_e_st_eq_landau_n_rt_rp_r_c_realc y) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isinv l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t1 x rx y ry i))))). Time Defined. (* constant 6831 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isrlc1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is r (l_e_st_eq_landau_n_rt_rp_r_c_rlofc (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_reali r))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_isst1 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t1 r). Time Defined. (* constant 6832 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isrlc2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_c_rlofc (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_reali r)) r). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_isst2 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t1 r). Time Defined. (* constant 6833 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscrl1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_rlofc x rx)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x) => l_e_ists1 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t1 x rx)). Time Defined. (* constant 6834 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscrl2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_rlofc x rx)) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (rx:l_e_st_eq_landau_n_rt_rp_r_c_realc x) => l_e_ists2 l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t1 x rx)). Time Defined. (* constant 6835 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_cofn : (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_complex). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_rlofnt n)). Time Defined. (* constant 6836 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnec : (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is n m), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_cofn m)))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is n m) => l_e_st_eq_landau_n_rt_rp_r_c_isrlec (l_e_st_eq_landau_n_rt_rp_r_rlofnt n) (l_e_st_eq_landau_n_rt_rp_r_rlofnt m) (l_e_st_eq_landau_n_rt_rp_r_isnterl n m i)))). Time Defined. (* constant 6837 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnic : (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_cofn m)), l_e_st_eq_landau_n_is n m))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_cofn m)) => l_e_st_eq_landau_n_rt_rp_r_isntirl n m (l_e_st_eq_landau_n_rt_rp_r_c_isrlic (l_e_st_eq_landau_n_rt_rp_r_rlofnt n) (l_e_st_eq_landau_n_rt_rp_r_rlofnt m) i)))). Time Defined. (* constant 6838 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v10_t2 : l_e_injective l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t). exact (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_nat) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofn t) (l_e_st_eq_landau_n_rt_rp_r_c_cofn u)) => l_e_st_eq_landau_n_rt_rp_r_c_isnic t u v))). Time Defined. (* constant 6839 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_natc : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), Prop). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_image l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) x). Time Defined. (* constant 6840 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_nati : (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_natc (l_e_st_eq_landau_n_rt_rp_r_c_cofn n)). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_imagei l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) n). Time Defined. (* constant 6841 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_nofc : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), l_e_st_eq_landau_n_nat)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => l_e_soft l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t2 x nx)). Time Defined. (* constant 6842 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscen : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (ny:l_e_st_eq_landau_n_rt_rp_r_c_natc y), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofc x nx) (l_e_st_eq_landau_n_rt_rp_r_c_nofc y ny)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (ny:l_e_st_eq_landau_n_rt_rp_r_c_natc y) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isinv l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t2 x nx y ny i))))). Time Defined. (* constant 6843 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscin : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (ny:l_e_st_eq_landau_n_rt_rp_r_c_natc y), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofc x nx) (l_e_st_eq_landau_n_rt_rp_r_c_nofc y ny)), l_e_st_eq_landau_n_rt_rp_r_c_is x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (ny:l_e_st_eq_landau_n_rt_rp_r_c_natc y) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofc x nx) (l_e_st_eq_landau_n_rt_rp_r_c_nofc y ny)) => l_e_isinve l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t2 x nx y ny i))))). Time Defined. (* constant 6844 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnc1 : (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is n (l_e_st_eq_landau_n_rt_rp_r_c_nofc (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_nati n))). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_isst1 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t2 n). Time Defined. (* constant 6845 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnc2 : (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofc (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_nati n)) n). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_isst2 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t2 n). Time Defined. (* constant 6846 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscn1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_cofn (l_e_st_eq_landau_n_rt_rp_r_c_nofc x nx)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => l_e_ists1 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t2 x nx)). Time Defined. (* constant 6847 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscn2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofn (l_e_st_eq_landau_n_rt_rp_r_c_nofc x nx)) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => l_e_ists2 l_e_st_eq_landau_n_nat l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_cofn t) l_e_st_eq_landau_n_rt_rp_r_c_v10_t2 x nx)). Time Defined. (* constant 6848 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_natt : Type. exact (l_e_ot l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t)). Time Defined. (* constant 6849 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_cofnt : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_st_eq_landau_n_rt_rp_r_c_cx). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_in l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) nt). Time Defined. (* constant 6850 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_natti : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_st_eq_landau_n_rt_rp_r_c_natc (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt)). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_inp l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) nt). Time Defined. (* constant 6851 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isntec : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt mt), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_cofnt mt)))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt mt) => l_e_isini l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) nt mt i))). Time Defined. (* constant 6852 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isntic : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_cofnt mt)), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt mt))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_cofnt mt)) => l_e_isine l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) nt mt i))). Time Defined. (* constant 6853 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ntofc : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), l_e_st_eq_landau_n_rt_rp_r_c_natt)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => l_e_out l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) x nx)). Time Defined. (* constant 6854 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscent : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (ny:l_e_st_eq_landau_n_rt_rp_r_c_natc y), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc x nx) (l_e_st_eq_landau_n_rt_rp_r_c_ntofc y ny)))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (ny:l_e_st_eq_landau_n_rt_rp_r_c_natc y) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is x y) => l_e_isouti l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) x nx y ny i))))). Time Defined. (* constant 6855 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscint : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), (forall (y:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (ny:l_e_st_eq_landau_n_rt_rp_r_c_natc y), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc x nx) (l_e_st_eq_landau_n_rt_rp_r_c_ntofc y ny)), l_e_st_eq_landau_n_rt_rp_r_c_is x y))))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => (fun (y:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (ny:l_e_st_eq_landau_n_rt_rp_r_c_natc y) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc x nx) (l_e_st_eq_landau_n_rt_rp_r_c_ntofc y ny)) => l_e_isoute l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) x nx y ny i))))). Time Defined. (* constant 6856 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isntc1 : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_isoutin l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) nt). Time Defined. (* constant 6857 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isntc2 : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt)) nt). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_natt nt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt)) (l_e_st_eq_landau_n_rt_rp_r_c_isntc1 nt)). Time Defined. (* constant 6858 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscnt1 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_cofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc x nx)))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => l_e_isinout l_e_st_eq_landau_n_rt_rp_r_c_cx (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_cx) => l_e_st_eq_landau_n_rt_rp_r_c_natc t) x nx)). Time Defined. (* constant 6859 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_iscnt2 : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc x nx)) x)). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_cx) => (fun (nx:l_e_st_eq_landau_n_rt_rp_r_c_natc x) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_cofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc x nx)) (l_e_st_eq_landau_n_rt_rp_r_c_iscnt1 x nx))). Time Defined. (* constant 6860 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ntofn : (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_natt). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_ntofc (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_nati n)). Time Defined. (* constant 6861 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnent : (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (i:l_e_st_eq_landau_n_is n m), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n) (l_e_st_eq_landau_n_rt_rp_r_c_ntofn m)))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (i:l_e_st_eq_landau_n_is n m) => l_e_st_eq_landau_n_rt_rp_r_c_iscent (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_nati n) (l_e_st_eq_landau_n_rt_rp_r_c_cofn m) (l_e_st_eq_landau_n_rt_rp_r_c_nati m) (l_e_st_eq_landau_n_rt_rp_r_c_isnec n m i)))). Time Defined. (* constant 6862 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnint : (forall (n:l_e_st_eq_landau_n_nat), (forall (m:l_e_st_eq_landau_n_nat), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n) (l_e_st_eq_landau_n_rt_rp_r_c_ntofn m)), l_e_st_eq_landau_n_is n m))). exact (fun (n:l_e_st_eq_landau_n_nat) => (fun (m:l_e_st_eq_landau_n_nat) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n) (l_e_st_eq_landau_n_rt_rp_r_c_ntofn m)) => l_e_st_eq_landau_n_rt_rp_r_c_isnic n m (l_e_st_eq_landau_n_rt_rp_r_c_iscint (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_nati n) (l_e_st_eq_landau_n_rt_rp_r_c_cofn m) (l_e_st_eq_landau_n_rt_rp_r_c_nati m) i)))). Time Defined. (* constant 6863 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_nofnt : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_st_eq_landau_n_nat). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_st_eq_landau_n_rt_rp_r_c_nofc (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt)). Time Defined. (* constant 6864 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnter : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt mt), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_nofnt mt)))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt mt) => l_e_st_eq_landau_n_rt_rp_r_c_iscen (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt) (l_e_st_eq_landau_n_rt_rp_r_c_cofnt mt) (l_e_st_eq_landau_n_rt_rp_r_c_natti mt) (l_e_st_eq_landau_n_rt_rp_r_c_isntec nt mt i)))). Time Defined. (* constant 6865 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isntin : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_nofnt mt)), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt mt))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (i:l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_nofnt mt)) => l_e_st_eq_landau_n_rt_rp_r_c_isntic nt mt (l_e_st_eq_landau_n_rt_rp_r_c_iscin (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt) (l_e_st_eq_landau_n_rt_rp_r_c_cofnt mt) (l_e_st_eq_landau_n_rt_rp_r_c_natti mt) i)))). Time Defined. (* constant 6866 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v10_t3 : (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_cofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n))). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_iscnt1 (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_nati n)). Time Defined. (* constant 6867 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnnt1 : (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is n (l_e_st_eq_landau_n_rt_rp_r_c_nofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n))). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_tris l_e_st_eq_landau_n_nat n (l_e_st_eq_landau_n_rt_rp_r_c_nofc (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_nati n)) (l_e_st_eq_landau_n_rt_rp_r_c_nofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n)) (l_e_st_eq_landau_n_rt_rp_r_c_isnc1 n) (l_e_st_eq_landau_n_rt_rp_r_c_iscen (l_e_st_eq_landau_n_rt_rp_r_c_cofn n) (l_e_st_eq_landau_n_rt_rp_r_c_nati n) (l_e_st_eq_landau_n_rt_rp_r_c_cofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n)) (l_e_st_eq_landau_n_rt_rp_r_c_natti (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n)) (l_e_st_eq_landau_n_rt_rp_r_c_v10_t3 n))). Time Defined. (* constant 6868 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isnnt2 : (forall (n:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n)) n). exact (fun (n:l_e_st_eq_landau_n_nat) => l_e_symis l_e_st_eq_landau_n_nat n (l_e_st_eq_landau_n_rt_rp_r_c_nofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n)) (l_e_st_eq_landau_n_rt_rp_r_c_isnnt1 n)). Time Defined. (* constant 6869 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_v10_t4 : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_cofn (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_st_eq_landau_n_rt_rp_r_c_iscn1 (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt)). Time Defined. (* constant 6870 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isntn1 : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_natt nt (l_e_st_eq_landau_n_rt_rp_r_c_ntofc (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt)) (l_e_st_eq_landau_n_rt_rp_r_c_ntofn (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) (l_e_st_eq_landau_n_rt_rp_r_c_isntc1 nt) (l_e_st_eq_landau_n_rt_rp_r_c_iscent (l_e_st_eq_landau_n_rt_rp_r_c_cofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_natti nt) (l_e_st_eq_landau_n_rt_rp_r_c_cofn (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) (l_e_st_eq_landau_n_rt_rp_r_c_nati (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) (l_e_st_eq_landau_n_rt_rp_r_c_v10_t4 nt))). Time Defined. (* constant 6871 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_isntn2 : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) nt). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_natt nt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) (l_e_st_eq_landau_n_rt_rp_r_c_isntn1 nt)). Time Defined. (* constant 6872 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_1t : l_e_st_eq_landau_n_rt_rp_r_c_natt. exact (l_e_st_eq_landau_n_rt_rp_r_c_ntofn l_e_st_eq_landau_n_1). Time Defined. (* constant 6873 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_suct : (forall (t:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_st_eq_landau_n_rt_rp_r_c_natt). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_st_eq_landau_n_rt_rp_r_c_ntofn (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt t))). Time Defined. (* constant 6874 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10299_t1 : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) l_e_st_eq_landau_n_rt_rp_r_c_1t), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) l_e_st_eq_landau_n_1)). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) l_e_st_eq_landau_n_rt_rp_r_c_1t) => l_e_st_eq_landau_n_rt_rp_r_c_isnint (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) l_e_st_eq_landau_n_1 i)). Time Defined. (* constant 6875 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz299a : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_not (l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) l_e_st_eq_landau_n_rt_rp_r_c_1t)). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_imp_th3 (l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) l_e_st_eq_landau_n_rt_rp_r_c_1t) (l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) l_e_st_eq_landau_n_1) (l_e_st_eq_landau_n_ax3 (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) (fun (t:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) l_e_st_eq_landau_n_rt_rp_r_c_1t) => l_e_st_eq_landau_n_rt_rp_r_c_10299_t1 nt t)). Time Defined. (* constant 6876 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ax3t : (forall (t:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_not (l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct t) l_e_st_eq_landau_n_rt_rp_r_c_1t)). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_st_eq_landau_n_rt_rp_r_c_satz299a t). Time Defined. (* constant 6877 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10299_t2 : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) (l_e_st_eq_landau_n_rt_rp_r_c_suct mt)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt mt))))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) (l_e_st_eq_landau_n_rt_rp_r_c_suct mt)) => l_e_st_eq_landau_n_rt_rp_r_c_isnint (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt)) (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt mt)) i))). Time Defined. (* constant 6878 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10299_t3 : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) (l_e_st_eq_landau_n_rt_rp_r_c_suct mt)), l_e_st_eq_landau_n_is (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_nofnt mt)))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) (l_e_st_eq_landau_n_rt_rp_r_c_suct mt)) => l_e_st_eq_landau_n_ax4 (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt) (l_e_st_eq_landau_n_rt_rp_r_c_nofnt mt) (l_e_st_eq_landau_n_rt_rp_r_c_10299_t2 nt mt i)))). Time Defined. (* constant 6879 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz299b : (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) (l_e_st_eq_landau_n_rt_rp_r_c_suct mt)), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt nt mt))). exact (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (mt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (i:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct nt) (l_e_st_eq_landau_n_rt_rp_r_c_suct mt)) => l_e_st_eq_landau_n_rt_rp_r_c_isntin nt mt (l_e_st_eq_landau_n_rt_rp_r_c_10299_t3 nt mt i)))). Time Defined. (* constant 6880 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ax4t : (forall (t:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (v:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct t) (l_e_st_eq_landau_n_rt_rp_r_c_suct u)), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt t u))). exact (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (v:l_e_is l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct t) (l_e_st_eq_landau_n_rt_rp_r_c_suct u)) => l_e_st_eq_landau_n_rt_rp_r_c_satz299b t u v))). Time Defined. (* constant 6881 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_cond1t : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt l_e_st_eq_landau_n_rt_rp_r_c_1t s). Time Defined. (* constant 6882 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_cond2t : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), Prop). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_all l_e_st_eq_landau_n_rt_rp_r_c_natt (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_imp (l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt t s) (l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_suct t) s))). Time Defined. (* constant 6883 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s), (forall (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s), (forall (n:l_e_st_eq_landau_n_nat), Prop)))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s) => (fun (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s) => (fun (n:l_e_st_eq_landau_n_nat) => l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n) s)))). Time Defined. (* constant 6884 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10299_t4 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s), (forall (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s), l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 l_e_st_eq_landau_n_1))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s) => (fun (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s) => c1))). Time Defined. (* constant 6885 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10299_t5 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s), (forall (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 n), l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn (l_e_st_eq_landau_n_suc (l_e_st_eq_landau_n_rt_rp_r_c_nofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n)))) s))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s) => (fun (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 n) => c2 (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n) p))))). Time Defined. (* constant 6886 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10299_t6 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s), (forall (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s), (forall (n:l_e_st_eq_landau_n_nat), (forall (p:l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 n), l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 (l_e_st_eq_landau_n_suc n)))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s) => (fun (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s) => (fun (n:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 n) => l_e_isp l_e_st_eq_landau_n_nat (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn (l_e_st_eq_landau_n_suc t)) s) (l_e_st_eq_landau_n_rt_rp_r_c_nofnt (l_e_st_eq_landau_n_rt_rp_r_c_ntofn n)) n (l_e_st_eq_landau_n_rt_rp_r_c_10299_t5 s c1 c2 n p) (l_e_st_eq_landau_n_rt_rp_r_c_isnnt2 n)))))). Time Defined. (* constant 6887 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10299_t7 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s), (forall (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s), (forall (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt))))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s) => (fun (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s) => (fun (nt:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_st_eq_landau_n_induction (fun (t:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 t) (l_e_st_eq_landau_n_rt_rp_r_c_10299_t4 s c1 c2) (fun (t:l_e_st_eq_landau_n_nat) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_10299_prop1 s c1 c2 t) => l_e_st_eq_landau_n_rt_rp_r_c_10299_t6 s c1 c2 t u)) (l_e_st_eq_landau_n_rt_rp_r_c_nofnt nt))))). Time Defined. (* constant 6888 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz299c : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s), (forall (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s), (forall (t:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt t s)))). exact (fun (s:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (c1:l_e_st_eq_landau_n_rt_rp_r_c_cond1t s) => (fun (c2:l_e_st_eq_landau_n_rt_rp_r_c_cond2t s) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_isp l_e_st_eq_landau_n_rt_rp_r_c_natt (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_natt) => l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt u s) (l_e_st_eq_landau_n_rt_rp_r_c_ntofn (l_e_st_eq_landau_n_rt_rp_r_c_nofnt t)) t (l_e_st_eq_landau_n_rt_rp_r_c_10299_t7 s c1 c2 t) (l_e_st_eq_landau_n_rt_rp_r_c_isntn2 t))))). Time Defined. (* constant 6889 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ax5t : (forall (t:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cond1t t), (forall (v:l_e_st_eq_landau_n_rt_rp_r_c_cond2t t), (forall (w:l_e_st_eq_landau_n_rt_rp_r_c_natt), l_e_st_esti l_e_st_eq_landau_n_rt_rp_r_c_natt w t)))). exact (fun (t:l_e_st_set l_e_st_eq_landau_n_rt_rp_r_c_natt) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_c_cond1t t) => (fun (v:l_e_st_eq_landau_n_rt_rp_r_c_cond2t t) => l_e_st_eq_landau_n_rt_rp_r_c_satz299c t u v))). Time Defined. (* constant 6890 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_ic : l_e_st_eq_landau_n_rt_rp_r_c_complex. exact (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl). Time Defined. (* constant 6891 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10300_t1 : l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_ic l_e_st_eq_landau_n_rt_rp_r_c_ic) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0))). exact (l_e_st_eq_landau_n_rt_rp_r_c_tsis12a l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl). Time Defined. (* constant 6892 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10300_t2 : l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl). exact (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl01 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ism0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl) l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_st_eq_landau_n_rt_rp_r_satz195 l_e_st_eq_landau_n_rt_rp_r_1rl))). Time Defined. (* constant 6893 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10300_t3 : l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0)) l_e_st_eq_landau_n_rt_rp_r_0. exact (l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl01 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_ts02 l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))). Time Defined. (* constant 6894 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10300_t4 : l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)). exact (l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0)) l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_c_10300_t2 l_e_st_eq_landau_n_rt_rp_r_c_10300_t3). Time Defined. (* constant 6895 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10300_t5 : l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c). exact (l_e_st_eq_landau_n_rt_rp_r_c_satz298j l_e_st_eq_landau_n_rt_rp_r_1rl). Time Defined. (* constant 6896 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz2300 : l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_ic l_e_st_eq_landau_n_rt_rp_r_c_ic) (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c). exact (l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts l_e_st_eq_landau_n_rt_rp_r_c_ic l_e_st_eq_landau_n_rt_rp_r_c_ic) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_1rl l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_m0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_c_m0 l_e_st_eq_landau_n_rt_rp_r_c_1c) l_e_st_eq_landau_n_rt_rp_r_c_10300_t1 l_e_st_eq_landau_n_rt_rp_r_c_10300_t4 l_e_st_eq_landau_n_rt_rp_r_c_10300_t5). Time Defined. (* constant 6897 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t1 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_tsis12a s l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)). Time Defined. (* constant 6898 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t2 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)) l_e_st_eq_landau_n_rt_rp_r_0)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl01 (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_m0 (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_ts02 s l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_satz176b (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))))). Time Defined. (* constant 6899 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t3 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_is (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) s)). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_real (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_1rl) s (l_e_st_eq_landau_n_rt_rp_r_pl02 (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts01 l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_satz195 s))). Time Defined. (* constant 6900 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t4 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)) l_e_st_eq_landau_n_rt_rp_r_0 (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0)) s (l_e_st_eq_landau_n_rt_rp_r_c_10301_t2 r s) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t3 r s))). Time Defined. (* constant 6901 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t5 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic) (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_mn (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_1rl)) (l_e_st_eq_landau_n_rt_rp_r_pl (l_e_st_eq_landau_n_rt_rp_r_ts s l_e_st_eq_landau_n_rt_rp_r_1rl) (l_e_st_eq_landau_n_rt_rp_r_ts l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0))) (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 s) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t1 r s) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t4 r s))). Time Defined. (* constant 6902 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t6 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_ispl2 (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic) (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 s) (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t5 r s))). Time Defined. (* constant 6903 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t7 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 s)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_plis12a r l_e_st_eq_landau_n_rt_rp_r_0 l_e_st_eq_landau_n_rt_rp_r_0 s)). Time Defined. (* constant 6904 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t8 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_c_isrecx12 (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_0) r (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 s) s (l_e_st_eq_landau_n_rt_rp_r_pl02 r l_e_st_eq_landau_n_rt_rp_r_0 (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)) (l_e_st_eq_landau_n_rt_rp_r_pl01 l_e_st_eq_landau_n_rt_rp_r_0 s (l_e_refis l_e_st_eq_landau_n_rt_rp_r_real l_e_st_eq_landau_n_rt_rp_r_0)))). Time Defined. (* constant 6905 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz301a : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pli r s))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_pli l_e_st_eq_landau_n_rt_rp_r_0 s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_pl r l_e_st_eq_landau_n_rt_rp_r_0) (l_e_st_eq_landau_n_rt_rp_r_pl l_e_st_eq_landau_n_rt_rp_r_0 s)) (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t6 r s) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t7 r s) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t8 r s))). Time Defined. (* constant 6906 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz301b : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) (l_e_st_eq_landau_n_rt_rp_r_c_satz301a r s))). Time Defined. (* constant 6907 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz301c : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is x (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_c_ic))). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_tris l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_pli (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_ispli x) (l_e_st_eq_landau_n_rt_rp_r_c_satz301b (l_e_st_eq_landau_n_rt_rp_r_c_re x) (l_e_st_eq_landau_n_rt_rp_r_c_im x))). Time Defined. (* constant 6908 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz301d : (forall (x:l_e_st_eq_landau_n_rt_rp_r_c_complex), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_c_ic)) x). exact (fun (x:l_e_st_eq_landau_n_rt_rp_r_c_complex) => l_e_symis l_e_st_eq_landau_n_rt_rp_r_c_cx x (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_re x)) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl (l_e_st_eq_landau_n_rt_rp_r_c_im x)) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_satz301c x)). Time Defined. (* constant 6909 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_10301_t9 : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl u) l_e_st_eq_landau_n_rt_rp_r_c_ic))), l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) (l_e_st_eq_landau_n_rt_rp_r_c_pli t u)))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl u) l_e_st_eq_landau_n_rt_rp_r_c_ic))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_c_cx (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl u) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pli t u) (l_e_st_eq_landau_n_rt_rp_r_c_satz301b r s) i (l_e_st_eq_landau_n_rt_rp_r_c_satz301a t u)))))). Time Defined. (* constant 6910 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz301e : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl u) l_e_st_eq_landau_n_rt_rp_r_c_ic))), l_e_st_eq_landau_n_rt_rp_r_is r t))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl u) l_e_st_eq_landau_n_rt_rp_r_c_ic))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real r (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_re (l_e_st_eq_landau_n_rt_rp_r_c_pli t u)) t (l_e_st_eq_landau_n_rt_rp_r_c_isre r s) (l_e_st_eq_landau_n_rt_rp_r_c_iscere (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) (l_e_st_eq_landau_n_rt_rp_r_c_pli t u) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t9 r s t u i)) (l_e_st_eq_landau_n_rt_rp_r_c_reis t u)))))). Time Defined. (* constant 6911 *) Definition l_e_st_eq_landau_n_rt_rp_r_c_satz301f : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), (forall (s:l_e_st_eq_landau_n_rt_rp_r_real), (forall (t:l_e_st_eq_landau_n_rt_rp_r_real), (forall (u:l_e_st_eq_landau_n_rt_rp_r_real), (forall (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl u) l_e_st_eq_landau_n_rt_rp_r_c_ic))), l_e_st_eq_landau_n_rt_rp_r_is s u))))). exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (s:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (t:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (u:l_e_st_eq_landau_n_rt_rp_r_real) => (fun (i:l_e_st_eq_landau_n_rt_rp_r_c_is (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl r) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl s) l_e_st_eq_landau_n_rt_rp_r_c_ic)) (l_e_st_eq_landau_n_rt_rp_r_c_pl (l_e_st_eq_landau_n_rt_rp_r_c_cofrl t) (l_e_st_eq_landau_n_rt_rp_r_c_ts (l_e_st_eq_landau_n_rt_rp_r_c_cofrl u) l_e_st_eq_landau_n_rt_rp_r_c_ic))) => l_e_tr3is l_e_st_eq_landau_n_rt_rp_r_real s (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli r s)) (l_e_st_eq_landau_n_rt_rp_r_c_im (l_e_st_eq_landau_n_rt_rp_r_c_pli t u)) u (l_e_st_eq_landau_n_rt_rp_r_c_isim r s) (l_e_st_eq_landau_n_rt_rp_r_c_isceim (l_e_st_eq_landau_n_rt_rp_r_c_pli r s) (l_e_st_eq_landau_n_rt_rp_r_c_pli t u) (l_e_st_eq_landau_n_rt_rp_r_c_10301_t9 r s t u i)) (l_e_st_eq_landau_n_rt_rp_r_c_imis t u)))))). Time Defined.